Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/thys/Bicategory/Bicategory.thy b/thys/Bicategory/Bicategory.thy
--- a/thys/Bicategory/Bicategory.thy
+++ b/thys/Bicategory/Bicategory.thy
@@ -1,2336 +1,2336 @@
(* Title: Bicategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
theory Bicategory
imports Prebicategory Category3.Subcategory Category3.DiscreteCategory
MonoidalCategory.MonoidalCategory
begin
section "Bicategories"
text \<open>
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.
\<close>
locale bicategory =
horizontal_composition V H src trg +
\<alpha>: natural_isomorphism VVV.comp V HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close> +
L: fully_faithful_functor V V L +
R: fully_faithful_functor V V R
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a" +
assumes unit_in_vhom: "obj a \<Longrightarrow> \<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
and iso_unit: "obj a \<Longrightarrow> iso \<i>[a]"
and pentagon: "\<lbrakk> ide f; ide g; ide h; ide k; src f = trg g; src g = trg h; src h = trg k \<rbrakk> \<Longrightarrow>
(f \<star> \<a>[g, h, k]) \<cdot> \<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k) = \<a>[f, g, h \<star> k] \<cdot> \<a>[f \<star> g, h, k]"
begin
(*
* TODO: the mapping \<i> is not currently assumed to be extensional.
* It might be best in the long run if it were.
*)
definition \<alpha>
where "\<alpha> \<mu>\<nu>\<tau> \<equiv> \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))"
lemma assoc_in_hom':
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "in_hhom \<a>[\<mu>, \<nu>, \<tau>] (src \<tau>) (trg \<mu>)"
and "\<guillemotleft>\<a>[\<mu>, \<nu>, \<tau>] : (dom \<mu> \<star> dom \<nu>) \<star> dom \<tau> \<Rightarrow> cod \<mu> \<star> cod \<nu> \<star> cod \<tau>\<guillemotright>"
proof -
show "\<guillemotleft>\<a>[\<mu>, \<nu>, \<tau>] : (dom \<mu> \<star> dom \<nu>) \<star> dom \<tau> \<Rightarrow> cod \<mu> \<star> cod \<nu> \<star> cod \<tau>\<guillemotright>"
proof -
have 1: "VVV.in_hom (\<mu>, \<nu>, \<tau>) (dom \<mu>, dom \<nu>, dom \<tau>) (cod \<mu>, cod \<nu>, cod \<tau>)"
- using assms VVV.in_hom_char VVV.arr_char VV.arr_char by auto
+ using assms VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
have "\<guillemotleft>\<a>[\<mu>, \<nu>, \<tau>] : HoHV (dom \<mu>, dom \<nu>, dom \<tau>) \<Rightarrow> HoVH (cod \<mu>, cod \<nu>, cod \<tau>)\<guillemotright>"
using 1 \<alpha>.preserves_hom by auto
moreover have "HoHV (dom \<mu>, dom \<nu>, dom \<tau>) = (dom \<mu> \<star> dom \<nu>) \<star> dom \<tau>"
- using 1 HoHV_def by (simp add: VVV.in_hom_char)
+ using 1 HoHV_def by (simp add: VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C)
moreover have "HoVH (cod \<mu>, cod \<nu>, cod \<tau>) = cod \<mu> \<star> cod \<nu> \<star> cod \<tau>"
- using 1 HoVH_def by (simp add: VVV.in_hom_char)
+ using 1 HoVH_def by (simp add: VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C)
ultimately show ?thesis by simp
qed
thus "in_hhom \<a>[\<mu>, \<nu>, \<tau>] (src \<tau>) (trg \<mu>)"
using assms src_cod trg_cod vconn_implies_hpar(1) vconn_implies_hpar(2) by auto
qed
lemma assoc_is_natural_1:
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "\<a>[\<mu>, \<nu>, \<tau>] = (\<mu> \<star> \<nu> \<star> \<tau>) \<cdot> \<a>[dom \<mu>, dom \<nu>, dom \<tau>]"
- using assms \<alpha>.is_natural_1 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char VV.arr_char VVV.dom_char
+ using assms \<alpha>.is_natural_1 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C
HoVH_def src_dom trg_dom
by simp
lemma assoc_is_natural_2:
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "\<a>[\<mu>, \<nu>, \<tau>] = \<a>[cod \<mu>, cod \<nu>, cod \<tau>] \<cdot> ((\<mu> \<star> \<nu>) \<star> \<tau>)"
- using assms \<alpha>.is_natural_2 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char VV.arr_char VVV.cod_char
+ using assms \<alpha>.is_natural_2 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
HoHV_def src_dom trg_dom
by simp
lemma assoc_naturality:
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "\<a>[cod \<mu>, cod \<nu>, cod \<tau>] \<cdot> ((\<mu> \<star> \<nu>) \<star> \<tau>) = (\<mu> \<star> \<nu> \<star> \<tau>) \<cdot> \<a>[dom \<mu>, dom \<nu>, dom \<tau>]"
- using assms \<alpha>.naturality VVV.arr_char VV.arr_char HoVH_def HoHV_def
- VVV.dom_char VVV.cod_char
+ using assms \<alpha>.naturality VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C HoVH_def HoHV_def
+ VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
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 \<a>[f, g, h] (src h) (trg f)"
and "\<guillemotleft>\<a>[f, g, h] : (dom f \<star> dom g) \<star> dom h \<Rightarrow> cod f \<star> cod g \<star> cod h\<guillemotright>"
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 \<a>[f, g, h]"
and "src \<a>[f, g, h] = src h" and "trg \<a>[f, g, h] = trg f"
and "dom \<a>[f, g, h] = (dom f \<star> dom g) \<star> dom h"
and "cod \<a>[f, g, h] = cod f \<star> cod g \<star> 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 \<a>[f, g, h]"
- using assms \<alpha>.components_are_iso [of "(f, g, h)"] VVV.ide_char VVV.arr_char VV.arr_char
+ using assms \<alpha>.components_are_iso [of "(f, g, h)"] VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
end
subsection "Categories Induce Bicategories"
text \<open>
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 \<open>src\<close> and \<open>trg\<close>.
\<close>
(*
* 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 \<open>Collect arr\<close> null
using not_arr_null by (unfold_locales, blast)
abbreviation V
where "V \<equiv> 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 \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<cdot> snd \<mu>\<nu>\<close>
apply (unfold_locales)
- using H.VV.arr_char V.null_char ext
+ using H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C V.null_char ext
apply force
- using H.VV.arr_char V.null_char H.VV.dom_char H.VV.cod_char
+ using H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C V.null_char H.VV.dom_char\<^sub>S\<^sub>b\<^sub>C H.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
apply auto[3]
proof -
show "\<And>g f. H.VV.seq g f \<Longrightarrow>
fst (H.VV.comp g f) \<cdot> snd (H.VV.comp g f) = V (fst g \<cdot> snd g) (fst f \<cdot> snd f)"
proof -
have 0: "\<And>f. H.VV.arr f \<Longrightarrow> V.arr (fst f \<cdot> snd f)"
- using H.VV.arr_char by auto
+ using H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
have 1: "\<And>f g. V.seq g f \<Longrightarrow> V.ide f \<and> g = f"
using V.arr_char V.dom_char V.cod_char V.not_arr_null by force
have 2: "\<And>f g. H.VxV.seq g f \<Longrightarrow> H.VxV.ide f \<and> 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 \<and> f = g"
- using fg 2 H.VV.seq_char H.VV.ide_char by blast
+ using fg 2 H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C H.VV.ide_char\<^sub>S\<^sub>b\<^sub>C by blast
show "fst (H.VV.comp g f) \<cdot> snd (H.VV.comp g f) = V (fst g \<cdot> snd g) (fst f \<cdot> 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
+ using fg 0 1 3 H.VV.comp_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<a>
where "\<a> f g h \<equiv> f \<cdot> g \<cdot> h"
interpretation \<alpha>: natural_isomorphism H.VVV.comp V H.HoHV H.HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close>
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
+ using H.HoHV_def H.HoVH_def H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C
+ H.VV.dom_char\<^sub>S\<^sub>b\<^sub>C H.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C H.VV.cod_char\<^sub>S\<^sub>b\<^sub>C 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 \<i>
where "\<i> \<equiv> \<lambda>x. x"
proposition induces_bicategory:
shows "bicategory V C \<a> \<i> dom cod"
apply (unfold_locales, auto simp add: comp_assoc)
using comp_arr_dom by fastforce
end
subsection "Monoidal Categories induce Bicategories"
text \<open>
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-\<open>\<I>\<close> functors
as src and trg and \<open>\<iota>\<close> as the unit isomorphism from \<open>\<I> \<otimes> \<I>\<close> to \<open>\<I>\<close>.
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.
\<close>
context monoidal_category
begin
interpretation I: constant_functor C C \<I>
using \<iota>_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 "\<And>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 \<Longrightarrow> CC.comp g f = VV.comp g f"
- using VV.comp_char VV.arr_char CC.seq_char
+ using VV.comp_char VV.arr_char\<^sub>S\<^sub>b\<^sub>C CC.seq_char
by (elim CC.seqE seqE, simp)
moreover have "\<not> CC.seq g f \<Longrightarrow> CC.comp g f = VV.comp g f"
- using VV.seq_char VV.ext VV.null_char CC.ext
+ using VV.seq_char\<^sub>S\<^sub>b\<^sub>C 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 "\<And>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 \<Longrightarrow> CCC.comp g f = VVV.comp g f"
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
+ I.preserves_reflects_arr VV.seq_char\<^sub>S\<^sub>b\<^sub>C VVV.arrI\<^sub>S\<^sub>b\<^sub>C VVV.comp_simp VVV.seq_char\<^sub>S\<^sub>b\<^sub>C
trg_vcomp vseq_implies_hpar(1))
moreover have "\<not> CCC.seq g f \<Longrightarrow> CCC.comp g f = VVV.comp g f"
- using VVV.seq_char VVV.ext VVV.null_char CCC.ext
+ using VVV.seq_char\<^sub>S\<^sub>b\<^sub>C 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 \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<otimes> snd \<mu>\<nu>\<close>
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 \<alpha>: natural_isomorphism VVV.comp C H.HoHV H.HoVH \<alpha>
using \<alpha>.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 \<longleftrightarrow> a = \<I>"
using obj_def [of a] \<iota>_in_hom by fastforce
proposition induces_bicategory:
shows "bicategory C tensor (\<lambda>\<mu> \<nu> \<tau>. \<alpha> (\<mu>, \<nu>, \<tau>)) (\<lambda>_. \<iota>) I.map I.map"
using obj_char \<iota>_in_hom \<iota>_is_iso pentagon \<alpha>.is_extensional \<alpha>.is_natural_1 \<alpha>.is_natural_2
by unfold_locales simp_all
end
subsection "Prebicategories Extend to Bicategories"
text \<open>
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 \<open>L\<close> and \<open>R\<close> 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 \<open>\<ll>\<close> and \<open>\<rr>\<close> from the endofunctors
\<open>L\<close> and \<open>R\<close> to the identity functor. A consequence is that functors \<open>L\<close> and \<open>R\<close> are
endo-equivalences, hence fully faithful.
\<close>
context prebicategory_with_homs
begin
text \<open>
Once it is equipped with a particular choice of source and target for each arrow,
a prebicategory determines a horizontal composition.
\<close>
lemma induces_horizontal_composition:
shows "horizontal_composition V H src trg"
proof -
interpret H: "functor" VV.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
proof -
have "VV.comp = VoV.comp"
using composable_char\<^sub>P\<^sub>B\<^sub>H by meson
thus "functor VV.comp V (\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>)"
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 \<subseteq> 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 ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text \<open>
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.
\<close>
definition lunit ("\<l>[_]")
where "lunit f \<equiv> left_hom_with_unit.lunit V H \<a> \<i>[trg f] (trg f) f"
definition runit ("\<r>[_]")
where "runit f \<equiv> right_hom_with_unit.runit V H \<a> \<i>[src f] (src f) f"
lemma lunit_in_hom:
assumes "ide f"
shows "\<guillemotleft>\<l>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>" and "\<guillemotleft>\<l>[f] : trg f \<star> f \<Rightarrow> f\<guillemotright>"
proof -
interpret Left: subcategory V \<open>left (trg f)\<close>
using assms left_hom_is_subcategory by simp
interpret Left: left_hom_with_unit V H \<a> \<open>\<i>[trg f]\<close> \<open>trg f\<close>
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
+ using assms Left.ide_char\<^sub>S\<^sub>b\<^sub>C Left.arr_char\<^sub>S\<^sub>b\<^sub>C left_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
show 1: "\<guillemotleft>\<l>[f] : trg f \<star> f \<Rightarrow> f\<guillemotright>"
unfolding lunit_def
using assms 0 Left.lunit_char(1) Left.hom_char H\<^sub>L_def by auto
show "\<guillemotleft>\<l>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>"
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 "\<guillemotleft>\<r>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>" and "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>"
proof -
interpret Right: subcategory V \<open>right (src f)\<close>
using assms right_hom_is_subcategory weak_unit_self_composable by force
interpret Right: right_hom_with_unit V H \<a> \<open>\<i>[src f]\<close> \<open>src f\<close>
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
+ using assms Right.ide_char\<^sub>S\<^sub>b\<^sub>C Right.arr_char\<^sub>S\<^sub>b\<^sub>C right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
show 1: "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>"
unfolding runit_def
using assms 0 Right.runit_char(1) Right.hom_char H\<^sub>R_def by auto
show "\<guillemotleft>\<r>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>"
using 1 src_cod trg_cod src_in_sources trg_in_targets
by (metis arrI vconn_implies_hpar)
qed
text \<open>
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.
\<close>
lemma lunit_char:
assumes "ide f"
shows "\<guillemotleft>\<l>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>" and "\<guillemotleft>\<l>[f] : trg f \<star> f \<Rightarrow> f\<guillemotright>"
and "trg f \<star> \<l>[f] = (\<i>[trg f] \<star> f) \<cdot> inv \<a>[trg f, trg f, f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : trg f \<star> f \<Rightarrow> f\<guillemotright> \<and> trg f \<star> \<mu> = (\<i>[trg f] \<star> f) \<cdot> inv \<a>[trg f, trg f, f]"
proof -
let ?a = "src f" and ?b = "trg f"
interpret Left: subcategory V \<open>left ?b\<close>
using assms left_hom_is_subcategory weak_unit_self_composable by force
interpret Left: left_hom_with_unit V H \<a> \<open>\<i>[?b]\<close> ?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
+ using assms Left.ide_char\<^sub>S\<^sub>b\<^sub>C Left.arr_char\<^sub>S\<^sub>b\<^sub>C left_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
show "\<guillemotleft>\<l>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>"
using assms lunit_in_hom by simp
show A: "\<guillemotleft>\<l>[f] : trg f \<star> f \<Rightarrow> f\<guillemotright>"
using assms lunit_in_hom by simp
show B: "?b \<star> \<l>[f] = (\<i>[?b] \<star> f) \<cdot> inv \<a>[?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 "\<exists>!\<mu>. \<guillemotleft>\<mu> : trg f \<star> f \<Rightarrow> f\<guillemotright> \<and> trg f \<star> \<mu> = (\<i>[?b] \<star> f) \<cdot> inv \<a>[?b, ?b, f]"
proof -
have 1: "hom (trg f \<star> f) f = Left.hom (Left.L f) f"
proof
have 1: "Left.L f = ?b \<star> f"
using 0 H\<^sub>L_def by simp
show "Left.hom (Left.L f) f \<subseteq> hom (?b \<star> f) f"
using assms Left.hom_char [of "?b \<star> f" f] H\<^sub>L_def by simp
show "hom (?b \<star> f) f \<subseteq> 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 = "\<lambda>\<mu>. Left.in_hom \<mu> (Left.L f) f"
let ?P' = "\<lambda>\<mu>. \<guillemotleft>\<mu> : ?b \<star> f \<Rightarrow> f\<guillemotright>"
let ?Q = "\<lambda>\<mu>. Left.L \<mu> = (\<i>[?b] \<star> f) \<cdot> (inv \<a>[?b, ?b, f])"
let ?R = "\<lambda>\<mu>. ?b \<star> \<mu> = (\<i>[?b] \<star> f) \<cdot> (inv \<a>[?b, ?b, f])"
have 2: "?P = ?P'"
using 0 1 H\<^sub>L_def Left.hom_char by blast
moreover have "\<forall>\<mu>. ?P \<mu> \<longrightarrow> (?Q \<mu> \<longleftrightarrow> ?R \<mu>)"
using 2 Left.lunit_eqI H\<^sub>L_def by presburger
moreover have "(\<exists>!\<mu>. ?P \<mu> \<and> ?Q \<mu>)"
- using 0 2 A B Left.lunit_char(3) Left.ide_char Left.arr_char
+ using 0 2 A B Left.lunit_char(3) Left.ide_char Left.arr_char\<^sub>S\<^sub>b\<^sub>C
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 "\<guillemotleft>\<r>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>" and "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>"
and "\<r>[f] \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : f \<star> src f \<Rightarrow> f\<guillemotright> \<and> \<mu> \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
proof -
let ?a = "src f" and ?b = "trg f"
interpret Right: subcategory V \<open>right ?a\<close>
using assms right_hom_is_subcategory weak_unit_self_composable by force
interpret Right: right_hom_with_unit V H \<a> \<open>\<i>[?a]\<close> ?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
+ using assms Right.ide_char\<^sub>S\<^sub>b\<^sub>C Right.arr_char\<^sub>S\<^sub>b\<^sub>C right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
show "\<guillemotleft>\<r>[f] : src f \<rightarrow>\<^sub>W\<^sub>C trg f\<guillemotright>"
using assms runit_in_hom by simp
show A: "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>"
using assms runit_in_hom by simp
show B: "\<r>[f] \<star> ?a = (f \<star> \<i>[?a]) \<cdot> \<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 "\<exists>!\<mu>. \<guillemotleft>\<mu> : f \<star> src f \<Rightarrow> f\<guillemotright> \<and> \<mu> \<star> ?a = (f \<star> \<i>[?a]) \<cdot> \<a>[f, ?a, ?a]"
proof -
have 1: "hom (f \<star> ?a) f = Right.hom (Right.R f) f"
proof
have 1: "Right.R f = f \<star> ?a"
using 0 H\<^sub>R_def by simp
show "Right.hom (Right.R f) f \<subseteq> hom (f \<star> ?a) f"
using assms Right.hom_char [of "f \<star> ?a" f] H\<^sub>R_def by simp
show "hom (f \<star> ?a) f \<subseteq> 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 = "\<lambda>\<mu>. Right.in_hom \<mu> (Right.R f) f"
let ?P' = "\<lambda>\<mu>. \<guillemotleft>\<mu> : f \<star> ?a \<Rightarrow> f\<guillemotright>"
let ?Q = "\<lambda>\<mu>. Right.R \<mu> = (f \<star> \<i>[?a]) \<cdot> \<a>[f, ?a, ?a]"
let ?R = "\<lambda>\<mu>. \<mu> \<star> ?a = (f \<star> \<i>[?a]) \<cdot> \<a>[f, ?a, ?a]"
have 2: "?P = ?P'"
using 0 1 H\<^sub>R_def Right.hom_char by blast
moreover have "\<forall>\<mu>. ?P \<mu> \<longrightarrow> (?Q \<mu> \<longleftrightarrow> ?R \<mu>)"
using 2 Right.runit_eqI H\<^sub>R_def by presburger
moreover have "(\<exists>!\<mu>. ?P \<mu> \<and> ?Q \<mu>)"
- using 0 2 A B Right.runit_char(3) Right.ide_char Right.arr_char
+ using 0 2 A B Right.runit_char(3) Right.ide_char Right.arr_char\<^sub>S\<^sub>b\<^sub>C
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 "\<guillemotleft>\<mu> : trg f \<star> f \<Rightarrow> f\<guillemotright>"
and "trg f \<star> \<mu> = (\<i>[trg f] \<star> f) \<cdot> (inv \<a>[trg f, trg f, f])"
shows "\<mu> = \<l>[f]"
using assms lunit_char(2-4) by blast
lemma runit_eqI:
assumes "ide f" and "\<guillemotleft>\<mu> : f \<star> src f \<Rightarrow> f\<guillemotright>"
and "\<mu> \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
shows "\<mu> = \<r>[f]"
using assms runit_char(2-4) by blast
lemma iso_lunit:
assumes "ide f"
shows "iso \<l>[f]"
proof -
let ?b = "trg f"
interpret Left: subcategory V \<open>left ?b\<close>
using assms left_hom_is_subcategory weak_unit_self_composable by force
interpret Left: left_hom_with_unit V H \<a> \<open>\<i>[?b]\<close> ?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
+ using assms Left.ide_char\<^sub>S\<^sub>b\<^sub>C Left.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<r>[f]"
proof -
let ?a = "src f"
interpret Right: subcategory V \<open>right ?a\<close>
using assms right_hom_is_subcategory weak_unit_self_composable by force
interpret Right: right_hom_with_unit V H \<a> \<open>\<i>[?a]\<close> ?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
+ using assms Right.ide_char\<^sub>S\<^sub>b\<^sub>C Right.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu>"
shows "\<mu> \<cdot> \<l>[dom \<mu>] = \<l>[cod \<mu>] \<cdot> (trg \<mu> \<star> \<mu>)"
proof -
let ?a = "src \<mu>" and ?b = "trg \<mu>"
interpret Left: subcategory V \<open>left ?b\<close>
using assms obj_trg left_hom_is_subcategory weak_unit_self_composable by force
interpret Left: left_hom_with_unit V H \<a> \<open>\<i>[?b]\<close> ?b
using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto)
interpret Left.L: endofunctor \<open>Left ?b\<close> 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 \<mu> (dom \<mu>) (cod \<mu>)"
- using assms Left.hom_char Left.arr_char left_def composable_char\<^sub>P\<^sub>B\<^sub>H obj_trg by auto
+ using assms Left.hom_char Left.arr_char\<^sub>S\<^sub>b\<^sub>C left_def composable_char\<^sub>P\<^sub>B\<^sub>H obj_trg by auto
have 2: "Left.in_hom \<l>[Left.dom \<mu>] (?b \<star> dom \<mu>) (dom \<mu>)"
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
+ using assms 1 Left.in_hom_char\<^sub>S\<^sub>b\<^sub>C trg_dom Left.lunit_char(1) H\<^sub>L_def
+ Left.arr_char\<^sub>S\<^sub>b\<^sub>C Left.dom_char\<^sub>S\<^sub>b\<^sub>C Left.ide_dom
by force
have 3: "Left.in_hom \<l>[Left.cod \<mu>] (?b \<star> cod \<mu>) (cod \<mu>)"
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
+ using assms 1 Left.in_hom_char\<^sub>S\<^sub>b\<^sub>C trg_cod Left.lunit_char(1) H\<^sub>L_def
+ Left.cod_char\<^sub>S\<^sub>b\<^sub>C Left.ide_cod
by force
have 4: "Left.in_hom (Left.L \<mu>) (?b \<star> dom \<mu>) (?b \<star> cod \<mu>)"
using 1 Left.L.preserves_hom [of \<mu> "dom \<mu>" "cod \<mu>"] H\<^sub>L_def by auto
show ?thesis
by (metis "1" "2" H\<^sub>L_def Left.comp_simp Left.in_homE Left.lunit_naturality
Left.seqI' lunit_def trg_cod trg_dom)
qed
lemma runit_naturality:
assumes "arr \<mu>"
shows "\<mu> \<cdot> \<r>[dom \<mu>] = \<r>[cod \<mu>] \<cdot> (\<mu> \<star> src \<mu>)"
proof -
let ?a = "src \<mu>" and ?b = "trg \<mu>"
interpret Right: subcategory V \<open>right ?a\<close>
using assms right_hom_is_subcategory weak_unit_self_composable by force
interpret Right: right_hom_with_unit V H \<a> \<open>\<i>[?a]\<close> ?a
using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto)
interpret Right.R: endofunctor \<open>Right ?a\<close> 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 \<mu> (dom \<mu>) (cod \<mu>)"
- using assms Right.hom_char Right.arr_char right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
+ using assms Right.hom_char Right.arr_char\<^sub>S\<^sub>b\<^sub>C right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
have 2: "Right.in_hom \<r>[Right.dom \<mu>] (dom \<mu> \<star> ?a) (dom \<mu>)"
unfolding runit_def
- using 1 Right.in_hom_char trg_dom Right.runit_char(1) [of "Right.dom \<mu>"] H\<^sub>R_def
- Right.arr_char Right.dom_char Right.ide_dom assms
+ using 1 Right.in_hom_char\<^sub>S\<^sub>b\<^sub>C trg_dom Right.runit_char(1) [of "Right.dom \<mu>"] H\<^sub>R_def
+ Right.arr_char\<^sub>S\<^sub>b\<^sub>C Right.dom_char\<^sub>S\<^sub>b\<^sub>C Right.ide_dom assms
by force
have 3: "\<r>[Right.cod \<mu>] \<in> Right.hom (cod \<mu> \<star> ?a) (cod \<mu>)"
unfolding runit_def
- using 1 Right.in_hom_char trg_cod Right.runit_char(1) [of "Right.cod \<mu>"] H\<^sub>R_def
- Right.cod_char Right.ide_cod assms
+ using 1 Right.in_hom_char\<^sub>S\<^sub>b\<^sub>C trg_cod Right.runit_char(1) [of "Right.cod \<mu>"] H\<^sub>R_def
+ Right.cod_char\<^sub>S\<^sub>b\<^sub>C Right.ide_cod assms
by force
have 4: "Right.R \<mu> \<in> Right.hom (dom \<mu> \<star> ?a) (cod \<mu> \<star> ?a)"
using 1 Right.R.preserves_hom [of \<mu> "dom \<mu>" "cod \<mu>"] H\<^sub>R_def by auto
show ?thesis
by (metis "1" "2" H\<^sub>R_def Right.comp_simp Right.in_homE Right.runit_naturality
Right.seqI' runit_def src_cod src_dom)
qed
interpretation L: endofunctor V L
using endofunctor_L by auto
interpretation \<ll>: transformation_by_components V V L map lunit
using lunit_in_hom lunit_naturality by unfold_locales auto
interpretation \<ll>: natural_isomorphism V V L map \<ll>.map
using iso_lunit by unfold_locales auto
lemma natural_isomorphism_\<ll>:
shows "natural_isomorphism V V L map \<ll>.map"
..
interpretation L: equivalence_functor V V L
using L.isomorphic_to_identity_is_equivalence \<ll>.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>[L f] = L \<l>[f]"
proof -
have "seq \<l>[f] (L \<l>[f])"
using assms lunit_char(2) L.preserves_hom by fastforce
moreover have "seq \<l>[f] \<l>[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 "\<l>[f]"] iso_lunit
iso_is_section section_is_mono monoE [of "\<l>[f]" "L \<l>[f]" "\<l>[L f]"]
by auto
qed
interpretation R: endofunctor V R
using endofunctor_R by auto
interpretation \<rr>: transformation_by_components V V R map runit
using runit_in_hom runit_naturality by unfold_locales auto
interpretation \<rr>: natural_isomorphism V V R map \<rr>.map
using iso_runit by unfold_locales auto
lemma natural_isomorphism_\<rr>:
shows "natural_isomorphism V V R map \<rr>.map"
..
interpretation R: equivalence_functor V V R
using R.isomorphic_to_identity_is_equivalence \<rr>.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>[R f] = R \<r>[f]"
proof -
have "seq \<r>[f] (R \<r>[f])"
using assms runit_char(2) R.preserves_hom by fastforce
moreover have "seq \<r>[f] \<r>[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 "\<r>[f]"] iso_runit
iso_is_section section_is_mono monoE [of "\<r>[f]" "R \<r>[f]" "\<r>[R f]"]
by auto
qed
definition \<alpha>
where "\<alpha> \<mu> \<nu> \<tau> \<equiv> if VVV.arr (\<mu>, \<nu>, \<tau>) then
(\<mu> \<star> \<nu> \<star> \<tau>) \<cdot> \<a>[dom \<mu>, dom \<nu>, dom \<tau>]
else null"
lemma \<alpha>_ide_simp [simp]:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "\<alpha> f g h = \<a>[f, g, h]"
proof -
have "\<alpha> f g h = (f \<star> g \<star> h) \<cdot> \<a>[dom f, dom g, dom h]"
- using assms \<alpha>_def VVV.arr_char [of "(f, g, h)"] by auto
+ using assms \<alpha>_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C [of "(f, g, h)"] by auto
also have "... = (f \<star> g \<star> h) \<cdot> \<a>[f, g, h]"
using assms by simp
also have "... = \<a>[f, g, h]"
- using assms \<alpha>_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
+ using assms \<alpha>_def assoc_in_hom\<^sub>A\<^sub>W\<^sub>C hcomp_in_hom\<^sub>P\<^sub>B\<^sub>H VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VoV.arr_char\<^sub>S\<^sub>b\<^sub>C
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 ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
lemma natural_isomorphism_\<alpha>:
shows "natural_isomorphism VVV.comp V HoHV HoVH
(\<lambda>\<mu>\<nu>\<tau>. \<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>)))"
proof -
interpret \<alpha>: transformation_by_components VVV.comp V HoHV HoVH
\<open>\<lambda>f. \<a>[fst f, fst (snd f), snd (snd f)]\<close>
proof
show 1: "\<And>x. VVV.ide x \<Longrightarrow> \<guillemotleft>\<a>[fst x, fst (snd x), snd (snd x)] : HoHV x \<Rightarrow> HoVH x\<guillemotright>"
proof -
fix x
assume x: "VVV.ide x"
show "\<guillemotleft>\<a>[fst x, fst (snd x), snd (snd x)] : HoHV x \<Rightarrow> HoVH x\<guillemotright>"
proof -
have "ide (fst x) \<and> ide (fst (snd x)) \<and> ide (snd (snd x)) \<and>
fst x \<star> fst (snd x) \<noteq> null \<and> fst (snd x) \<star> snd (snd x) \<noteq> null"
- using x VVV.ide_char VVV.arr_char VV.arr_char composable_char\<^sub>P\<^sub>B\<^sub>H by simp
+ using x VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C composable_char\<^sub>P\<^sub>B\<^sub>H by simp
hence "\<a>[fst x, fst (snd x), snd (snd x)]
\<in> hom ((fst x \<star> fst (snd x)) \<star> snd (snd x))
(fst x \<star> fst (snd x) \<star> 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 "\<And>f. VVV.arr f \<Longrightarrow>
\<a>[fst (VVV.cod f), fst (snd (VVV.cod f)), snd (snd (VVV.cod f))] \<cdot> HoHV f =
HoVH f \<cdot> \<a>[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
+ using assoc_naturality\<^sub>A\<^sub>W\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
composable_char\<^sub>P\<^sub>B\<^sub>H
by simp
qed
interpret \<alpha>: natural_isomorphism VVV.comp V HoHV HoVH \<alpha>.map
proof
fix f
assume f: "VVV.ide f"
show "iso (\<alpha>.map f)"
proof -
have "fst f \<star> fst (snd f) \<noteq> null \<and> fst (snd f) \<star> snd (snd f) \<noteq> null"
- using f VVV.ideD(1) VVV.arr_char [of f] VV.arr_char composable_char\<^sub>P\<^sub>B\<^sub>H by auto
+ using f VVV.ideD(1) VVV.arr_char\<^sub>S\<^sub>b\<^sub>C [of f] VV.arr_char\<^sub>S\<^sub>b\<^sub>C composable_char\<^sub>P\<^sub>B\<^sub>H by auto
thus ?thesis
- using f \<alpha>.map_simp_ide iso_assoc\<^sub>A\<^sub>W\<^sub>C VVV.ide_char VVV.arr_char by simp
+ using f \<alpha>.map_simp_ide iso_assoc\<^sub>A\<^sub>W\<^sub>C VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
qed
have "(\<lambda>\<mu>\<nu>\<tau>. \<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))) = \<alpha>.map"
proof
fix \<mu>\<nu>\<tau>
have "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> \<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>)) = \<alpha>.map \<mu>\<nu>\<tau>"
using \<alpha>_def \<alpha>.map_def by simp
moreover have "VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow>
\<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>)) = \<alpha>.map \<mu>\<nu>\<tau>"
proof -
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have "\<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>)) =
(fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>) \<star> snd (snd \<mu>\<nu>\<tau>)) \<cdot>
\<a>[dom (fst \<mu>\<nu>\<tau>), dom (fst (snd \<mu>\<nu>\<tau>)), dom (snd (snd \<mu>\<nu>\<tau>))]"
using \<mu>\<nu>\<tau> \<alpha>_def by simp
also have "... = \<a>[cod (fst \<mu>\<nu>\<tau>), cod (fst (snd \<mu>\<nu>\<tau>)), cod (snd (snd \<mu>\<nu>\<tau>))] \<cdot>
((fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>)) \<star> snd (snd \<mu>\<nu>\<tau>))"
- using \<mu>\<nu>\<tau> HoHV_def HoVH_def VVV.arr_char VV.arr_char assoc_naturality\<^sub>A\<^sub>W\<^sub>C
+ using \<mu>\<nu>\<tau> HoHV_def HoVH_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C assoc_naturality\<^sub>A\<^sub>W\<^sub>C
composable_char\<^sub>P\<^sub>B\<^sub>H
by simp
also have "... =
\<a>[fst (VVV.cod \<mu>\<nu>\<tau>), fst (snd (VVV.cod \<mu>\<nu>\<tau>)), snd (snd (VVV.cod \<mu>\<nu>\<tau>))] \<cdot>
((fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>)) \<star> snd (snd \<mu>\<nu>\<tau>))"
- using \<mu>\<nu>\<tau> VVV.arr_char VVV.cod_char VV.arr_char by simp
+ using \<mu>\<nu>\<tau> VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = \<alpha>.map \<mu>\<nu>\<tau>"
using \<mu>\<nu>\<tau> \<alpha>.map_def HoHV_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto
finally show ?thesis by blast
qed
ultimately show "\<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>)) = \<alpha>.map \<mu>\<nu>\<tau>" by blast
qed
thus ?thesis using \<alpha>.natural_isomorphism_axioms by simp
qed
proposition induces_bicategory:
shows "bicategory V H \<alpha> \<i> src trg"
proof -
interpret VxVxV: product_category V VxV.comp ..
interpret VoVoV: subcategory VxVxV.comp
\<open>\<lambda>\<tau>\<mu>\<nu>. arr (fst \<tau>\<mu>\<nu>) \<and> VV.arr (snd \<tau>\<mu>\<nu>) \<and>
src (fst \<tau>\<mu>\<nu>) = trg (fst (snd \<tau>\<mu>\<nu>))\<close>
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 \<alpha>: natural_isomorphism VVV.comp V HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<alpha> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close>
using natural_isomorphism_\<alpha> 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 \<alpha> \<i> src trg"
proof
show "\<And>a. obj a \<Longrightarrow> \<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
using obj_is_weak_unit unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by blast
show "\<And>a. obj a \<Longrightarrow> iso \<i>[a]"
using obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by blast
show "\<And>f g h k. \<lbrakk> ide f; ide g; ide h; ide k;
src f = trg g; src g = trg h; src h = trg k \<rbrakk> \<Longrightarrow>
(f \<star> \<alpha> g h k) \<cdot> \<alpha> f (g \<star> h) k \<cdot> (\<alpha> f g h \<star> k) =
\<alpha> f g (h \<star> k) \<cdot> \<alpha> (f \<star> 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 \<inter> targets g \<noteq> {}"
using f g fg src_in_sources [of f] trg_in_targets ideD(1) by auto
moreover have "sources g \<inter> targets h \<noteq> {}"
using g h gh src_in_sources [of g] trg_in_targets ideD(1) by auto
moreover have "sources h \<inter> targets k \<noteq> {}"
using h k hk src_in_sources [of h] trg_in_targets ideD(1) by auto
moreover have "\<alpha> f g h = \<a>[f, g, h] \<and> \<alpha> g h k = \<a>[g, h, k]"
using f g h k fg gh hk \<alpha>_ide_simp by simp
moreover have "\<alpha> f (g \<star> h) k = \<a>[f, g \<star> h, k] \<and> \<alpha> f g (h \<star> k) = \<a>[f, g, h \<star> k] \<and>
\<alpha> (f \<star> g) h k = \<a>[f \<star> g, h, k]"
using f g h k fg gh hk \<alpha>_ide_simp preserves_ide hcomp_in_hom\<^sub>P\<^sub>B\<^sub>H(1) by simp
ultimately show "(f \<star> \<alpha> g h k) \<cdot> \<alpha> f (g \<star> h) k \<cdot> (\<alpha> f g h \<star> k) =
\<alpha> f g (h \<star> k) \<cdot> \<alpha> (f \<star> g) h k"
using f g h k fg gh hk pentagon\<^sub>A\<^sub>W\<^sub>C [of f g h k] \<alpha>_ide_simp by presburger
qed
qed
qed
end
text \<open>
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 \<open>hom (a \<star> a) a\<close> for each weak unit \<open>a\<close>.
\<close>
context prebicategory
begin
interpretation prebicategory_with_homs V H \<a> some_src some_trg
using extends_to_prebicategory_with_homs by auto
interpretation prebicategory_with_units V H \<a> some_unit
using extends_to_prebicategory_with_units by auto
interpretation prebicategory_with_homs_and_units V H \<a> some_unit some_src some_trg ..
theorem extends_to_bicategory:
shows "bicategory V H \<alpha> some_unit some_src some_trg"
using induces_bicategory by simp
end
section "Bicategories as Prebicategories"
subsection "Bicategories are Prebicategories"
text \<open>
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.
\<close>
context bicategory
begin
(* TODO: Why does this get re-introduced? *)
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
interpretation \<alpha>': inverse_transformation VVV.comp V HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close> ..
abbreviation \<alpha>'
where "\<alpha>' \<equiv> \<alpha>'.map"
definition \<a>' ("\<a>\<^sup>-\<^sup>1[_, _, _]")
where "\<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] \<equiv> \<alpha>'.map (\<mu>, \<nu>, \<tau>)"
lemma assoc'_in_hom':
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "in_hhom \<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] (src \<tau>) (trg \<mu>)"
and "\<guillemotleft>\<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] : dom \<mu> \<star> dom \<nu> \<star> dom \<tau> \<Rightarrow> (cod \<mu> \<star> cod \<nu>) \<star> cod \<tau>\<guillemotright>"
proof -
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] : dom \<mu> \<star> dom \<nu> \<star> dom \<tau> \<Rightarrow> (cod \<mu> \<star> cod \<nu>) \<star> cod \<tau>\<guillemotright>"
proof -
have 1: "VVV.in_hom (\<mu>, \<nu>, \<tau>) (dom \<mu>, dom \<nu>, dom \<tau>) (cod \<mu>, cod \<nu>, cod \<tau>)"
- using assms VVV.in_hom_char VVV.arr_char VV.arr_char by auto
+ using assms VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
have "\<guillemotleft>\<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] : HoVH (dom \<mu>, dom \<nu>, dom \<tau>) \<Rightarrow> HoHV (cod \<mu>, cod \<nu>, cod \<tau>)\<guillemotright>"
using 1 \<a>'_def \<alpha>'.preserves_hom by auto
moreover have "HoVH (dom \<mu>, dom \<nu>, dom \<tau>) = dom \<mu> \<star> dom \<nu> \<star> dom \<tau>"
- using 1 HoVH_def by (simp add: VVV.in_hom_char)
+ using 1 HoVH_def by (simp add: VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C)
moreover have "HoHV (cod \<mu>, cod \<nu>, cod \<tau>) = (cod \<mu> \<star> cod \<nu>) \<star> cod \<tau>"
- using 1 HoHV_def by (simp add: VVV.in_hom_char)
+ using 1 HoHV_def by (simp add: VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C)
ultimately show ?thesis by simp
qed
thus "in_hhom \<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] (src \<tau>) (trg \<mu>)"
using assms vconn_implies_hpar(1) vconn_implies_hpar(2) by auto
qed
lemma assoc'_is_natural_1:
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "\<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] = ((\<mu> \<star> \<nu>) \<star> \<tau>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<mu>, dom \<nu>, dom \<tau>]"
- using assms \<alpha>'.is_natural_1 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char VV.arr_char
- VVV.dom_char HoHV_def src_dom trg_dom \<a>'_def
+ using assms \<alpha>'.is_natural_1 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ VVV.dom_char\<^sub>S\<^sub>b\<^sub>C HoHV_def src_dom trg_dom \<a>'_def
by simp
lemma assoc'_is_natural_2:
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "\<a>\<^sup>-\<^sup>1[\<mu>, \<nu>, \<tau>] = \<a>\<^sup>-\<^sup>1[cod \<mu>, cod \<nu>, cod \<tau>] \<cdot> (\<mu> \<star> \<nu> \<star> \<tau>)"
- using assms \<alpha>'.is_natural_2 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char VV.arr_char
- VVV.cod_char HoVH_def src_dom trg_dom \<a>'_def
+ using assms \<alpha>'.is_natural_2 [of "(\<mu>, \<nu>, \<tau>)"] VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ VVV.cod_char\<^sub>S\<^sub>b\<^sub>C HoVH_def src_dom trg_dom \<a>'_def
by simp
lemma assoc'_naturality:
assumes "arr \<mu>" and "arr \<nu>" and "arr \<tau>" and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "\<a>\<^sup>-\<^sup>1[cod \<mu>, cod \<nu>, cod \<tau>] \<cdot> (\<mu> \<star> \<nu> \<star> \<tau>) = ((\<mu> \<star> \<nu>) \<star> \<tau>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<mu>, dom \<nu>, dom \<tau>]"
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 \<a>\<^sup>-\<^sup>1[f, g, h] (src h) (trg f)"
and "\<guillemotleft>\<a>\<^sup>-\<^sup>1[f, g, h] : dom f \<star> dom g \<star> dom h \<Rightarrow> (cod f \<star> cod g) \<star> cod h\<guillemotright>"
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 \<a>\<^sup>-\<^sup>1[f, g, h]"
and "src \<a>\<^sup>-\<^sup>1[f, g, h] = src h" and "trg \<a>\<^sup>-\<^sup>1[f, g, h] = trg f"
and "dom \<a>\<^sup>-\<^sup>1[f, g, h] = dom f \<star> dom g \<star> dom h"
and "cod \<a>\<^sup>-\<^sup>1[f, g, h] = (cod f \<star> cod g) \<star> 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 "\<a>\<^sup>-\<^sup>1[f, g, h] = inv \<a>[f, g, h]"
- using assms VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char \<alpha>'.map_ide_simp
+ using assms VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<alpha>'.map_ide_simp
\<a>'_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 \<a>[f, g, h] \<a>\<^sup>-\<^sup>1[f, g, h]"
- using assms VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char \<alpha>'.map_ide_simp
+ using assms VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<alpha>'.map_ide_simp
\<alpha>'.inverts_components \<a>'_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 \<a>\<^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 "\<a>[f, g, h] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, h] = f \<star> g \<star> h"
and "\<a>\<^sup>-\<^sup>1[f, g, h] \<cdot> \<a>[f, g, h] = (f \<star> g) \<star> h"
using assms comp_arr_inv' comp_inv_arr' by auto
lemma unit_in_hom [intro, simp]:
assumes "obj a"
shows "\<guillemotleft>\<i>[a] : a \<rightarrow> a\<guillemotright>" and "\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
proof -
show "\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
using assms unit_in_vhom by simp
thus "\<guillemotleft>\<i>[a] : a \<rightarrow> a\<guillemotright>"
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 "\<nu> \<star> \<mu> \<noteq> null"
shows "src \<nu> = trg \<mu>"
- using assms H.is_extensional [of "(\<nu>, \<mu>)"] VV.arr_char by auto
+ using assms H.is_extensional [of "(\<nu>, \<mu>)"] VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
lemma obj_self_composable:
assumes "obj a"
shows "a \<star> a \<noteq> null"
and "isomorphic (a \<star> a) a"
apply (metis arr_dom_iff_arr assms in_homE not_arr_null unit_in_vhom)
by (meson assms iso_unit isomorphic_def unit_in_vhom)
lemma obj_is_weak_unit:
assumes "obj a"
shows "weak_unit a"
proof -
interpret Left_a: subcategory V \<open>left a\<close>
using assms left_hom_is_subcategory by force
interpret Right_a: subcategory V \<open>right a\<close>
using assms right_hom_is_subcategory by force
text \<open>
We know that \<open>H\<^sub>L a\<close> 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.
\<close>
interpret La: endofunctor \<open>Left a\<close> \<open>H\<^sub>L a\<close>
using assms obj_self_composable endofunctor_H\<^sub>L [of a] by force
interpret La: fully_faithful_functor \<open>Left a\<close> \<open>Left a\<close> \<open>H\<^sub>L a\<close>
proof
show "\<And>f f'. Left_a.par f f' \<Longrightarrow> H\<^sub>L a f = H\<^sub>L a f' \<Longrightarrow> f = f'"
proof -
fix \<mu> \<mu>'
assume par: "Left_a.par \<mu> \<mu>'"
assume eq: "H\<^sub>L a \<mu> = H\<^sub>L a \<mu>'"
have 1: "par \<mu> \<mu>'"
- using par Left_a.arr_char Left_a.dom_char Left_a.cod_char left_def
+ using par Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C Left_a.dom_char\<^sub>S\<^sub>b\<^sub>C Left_a.cod_char\<^sub>S\<^sub>b\<^sub>C left_def
composable_implies_arr null_agreement
by metis
moreover have "L \<mu> = L \<mu>'"
- using par eq H\<^sub>L_def Left_a.arr_char left_def preserves_arr
+ using par eq H\<^sub>L_def Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def preserves_arr
assms 1 seq_if_composable [of a \<mu>] not_arr_null seq_if_composable [of a \<mu>']
by auto
ultimately show "\<mu> = \<mu>'"
using L.is_faithful by blast
qed
show "\<And>f g \<mu>. \<lbrakk> Left_a.ide f; Left_a.ide g; Left_a.in_hom \<mu> (H\<^sub>L a f) (H\<^sub>L a g) \<rbrakk> \<Longrightarrow>
\<exists>\<nu>. Left_a.in_hom \<nu> f g \<and> H\<^sub>L a \<nu> = \<mu>"
proof -
fix f g \<mu>
assume f: "Left_a.ide f" and g: "Left_a.ide g"
and \<mu>: "Left_a.in_hom \<mu> (H\<^sub>L a f) (H\<^sub>L a g)"
have 1: "a = trg f \<and> a = trg g"
- using assms f g Left_a.ide_char Left_a.arr_char left_def seq_if_composable [of a f]
+ using assms f g Left_a.ide_char Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def seq_if_composable [of a f]
seq_if_composable [of a g]
by auto
show "\<exists>\<nu>. Left_a.in_hom \<nu> f g \<and> H\<^sub>L a \<nu> = \<mu>"
proof -
have 2: "\<exists>\<nu>. \<guillemotleft>\<nu> : f \<Rightarrow> g\<guillemotright> \<and> L \<nu> = \<mu>"
- using f g \<mu> 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
+ using f g \<mu> 1 Left_a.ide_char\<^sub>S\<^sub>b\<^sub>C H\<^sub>L_def L.preserves_reflects_arr Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C
+ Left_a.in_hom_char\<^sub>S\<^sub>b\<^sub>C L.is_full
by force
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu> : f \<Rightarrow> g\<guillemotright> \<and> L \<nu> = \<mu>"
using 2 by blast
have "Left_a.arr \<nu>"
- using \<nu> 1 trg_dom Left_a.arr_char left_def hseq_char' by fastforce
+ using \<nu> 1 trg_dom Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def hseq_char' by fastforce
moreover have "H\<^sub>L a \<nu> = \<mu>"
using \<nu> 1 trg_dom H\<^sub>L_def by auto
ultimately show ?thesis
using \<nu> Left_a.dom_simp Left_a.cod_simp by blast
qed
qed
qed
interpret Ra: endofunctor \<open>Right a\<close> \<open>H\<^sub>R a\<close>
using assms obj_self_composable endofunctor_H\<^sub>R [of a] by force
interpret Ra: fully_faithful_functor \<open>Right a\<close> \<open>Right a\<close> \<open>H\<^sub>R a\<close>
proof
show "\<And>f f'. Right_a.par f f' \<Longrightarrow> H\<^sub>R a f = H\<^sub>R a f' \<Longrightarrow> f = f'"
proof -
fix \<mu> \<mu>'
assume par: "Right_a.par \<mu> \<mu>'"
assume eq: "H\<^sub>R a \<mu> = H\<^sub>R a \<mu>'"
have 1: "par \<mu> \<mu>'"
- using par Right_a.arr_char Right_a.dom_char Right_a.cod_char right_def
+ using par Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C Right_a.dom_char\<^sub>S\<^sub>b\<^sub>C Right_a.cod_char\<^sub>S\<^sub>b\<^sub>C right_def
composable_implies_arr null_agreement
by metis
moreover have "R \<mu> = R \<mu>'"
- using par eq H\<^sub>R_def Right_a.arr_char right_def preserves_arr
+ using par eq H\<^sub>R_def Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def preserves_arr
assms 1 seq_if_composable [of \<mu> a] not_arr_null seq_if_composable [of \<mu>' a]
by auto
ultimately show "\<mu> = \<mu>'"
using R.is_faithful by blast
qed
show "\<And>f g \<mu>. \<lbrakk> Right_a.ide f; Right_a.ide g; Right_a.in_hom \<mu> (H\<^sub>R a f) (H\<^sub>R a g) \<rbrakk> \<Longrightarrow>
\<exists>\<nu>. Right_a.in_hom \<nu> f g \<and> H\<^sub>R a \<nu> = \<mu>"
proof -
fix f g \<mu>
assume f: "Right_a.ide f" and g: "Right_a.ide g"
and \<mu>: "Right_a.in_hom \<mu> (H\<^sub>R a f) (H\<^sub>R a g)"
have 1: "a = src f \<and> a = src g"
- using assms f g Right_a.ide_char Right_a.arr_char right_def seq_if_composable
+ using assms f g Right_a.ide_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def seq_if_composable
by auto
show "\<exists>\<nu>. Right_a.in_hom \<nu> f g \<and> H\<^sub>R a \<nu> = \<mu>"
proof -
have 2: "\<exists>\<nu>. \<guillemotleft>\<nu> : f \<Rightarrow> g\<guillemotright> \<and> R \<nu> = \<mu>"
- using f g \<mu> 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
+ using f g \<mu> 1 Right_a.ide_char\<^sub>S\<^sub>b\<^sub>C H\<^sub>R_def R.preserves_reflects_arr Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C
+ Right_a.in_hom_char\<^sub>S\<^sub>b\<^sub>C R.is_full
by force
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu> : f \<Rightarrow> g\<guillemotright> \<and> R \<nu> = \<mu>"
using 2 by blast
have "Right_a.arr \<nu>"
- using \<nu> 1 src_dom Right_a.arr_char right_def hseq_char' by fastforce
+ using \<nu> 1 src_dom Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def hseq_char' by fastforce
moreover have "H\<^sub>R a \<nu> = \<mu>"
using \<nu> 1 src_dom H\<^sub>R_def by auto
ultimately show ?thesis
using \<nu> Right_a.dom_simp Right_a.cod_simp by blast
qed
qed
qed
have "isomorphic (a \<star> a) a \<and> a \<star> a \<noteq> 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 \<mu>"
shows "src \<mu> \<in> sources \<mu>"
using assms obj_is_weak_unit R.preserves_arr hseq_char' by auto
lemma trg_in_targets:
assumes "arr \<mu>"
shows "trg \<mu> \<in> targets \<mu>"
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 \<star> f \<cong> a \<star> g"
shows "f \<cong> g"
proof -
have 0: "ide a"
using assms weak_unit_def by force
interpret Left_a: subcategory V \<open>left a\<close>
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 \<open>Left a\<close> \<open>Left a\<close> \<open>H\<^sub>L a\<close>
using assms weak_unit_def by fast
obtain \<phi> where \<phi>: "iso \<phi> \<and> \<guillemotleft>\<phi> : a \<star> f \<Rightarrow> a \<star> g\<guillemotright>"
using assms by blast
have 1: "Left_a.iso \<phi> \<and> Left_a.in_hom \<phi> (a \<star> f) (a \<star> g)"
- by (metis H\<^sub>L_def La.is_extensional La.preserves_arr Left_a.arrI Left_a.dom_closed
- Left_a.in_hom_char Left_a.iso_char Left_a.not_arr_null Left_a.null_char \<phi> assms(4)
+ by (metis H\<^sub>L_def La.is_extensional La.preserves_arr Left_a.dom_closed
+ Left_a.in_hom_char\<^sub>S\<^sub>b\<^sub>C Left_a.iso_char Left_a.not_arr_null Left_a.null_char \<phi> assms(4)
hom_connected(4) hseq_char' ide_char' in_homE isomorphic_implies_ide(2) left_def)
hence 2: "Left_a.ide (a \<star> f) \<and> Left_a.ide (a \<star> g)"
using Left_a.ide_dom [of \<phi>] Left_a.ide_cod [of \<phi>] Left_a.dom_simp Left_a.cod_simp
by auto
hence 3: "Left_a.ide f \<and> Left_a.ide g"
- by (metis Left_a.ideI Left_a.ide_def Left_a.null_char assms(2) assms(3) left_def)
+ by (metis Left_a.ideI\<^sub>S\<^sub>b\<^sub>C Left_a.ide_def Left_a.null_char assms(2) assms(3) left_def)
obtain \<psi> where \<psi>: "\<psi> \<in> Left_a.hom f g \<and> a \<star> \<psi> = \<phi>"
using assms 1 2 3 La.is_full [of g f \<phi>] H\<^sub>L_def by auto
have "Left_a.iso \<psi>"
using \<psi> 1 H\<^sub>L_def La.reflects_iso by auto
hence "iso \<psi> \<and> \<guillemotleft>\<psi> : f \<Rightarrow> g\<guillemotright>"
- using \<psi> Left_a.iso_char Left_a.in_hom_char by auto
+ using \<psi> Left_a.iso_char Left_a.in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis by auto
qed
lemma weak_unit_cancel_right:
assumes "weak_unit a" and "ide f" and "ide g"
and "f \<star> a \<cong> g \<star> a"
shows "f \<cong> g"
proof -
have 0: "ide a"
using assms weak_unit_def by force
interpret Right_a: subcategory V \<open>right a\<close>
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 \<open>Right a\<close> \<open>Right a\<close> \<open>H\<^sub>R a\<close>
using assms weak_unit_def by fast
obtain \<phi> where \<phi>: "iso \<phi> \<and> in_hom \<phi> (f \<star> a) (g \<star> a)"
using assms by blast
have 1: "Right_a.iso \<phi> \<and> \<phi> \<in> Right_a.hom (f \<star> a) (g \<star> a)"
proof
have "\<phi> \<star> a \<noteq> null"
by (metis \<phi> assms(1,4) hom_connected(3) ideD(1) in_homE isomorphic_implies_ide(2)
match_3 not_arr_null weak_unit_self_composable(3))
thus "\<phi> \<in> Right_a.hom (f \<star> a) (g \<star> a)"
using \<phi> Right_a.hom_char right_def by simp
thus "Right_a.iso \<phi>"
using \<phi> Right_a.iso_char by auto
qed
hence 2: "Right_a.ide (f \<star> a) \<and> Right_a.ide (g \<star> a)"
using Right_a.ide_dom [of \<phi>] Right_a.ide_cod [of \<phi>] Right_a.dom_simp Right_a.cod_simp
by auto
hence 3: "Right_a.ide f \<and> Right_a.ide g"
- using assms Right_a.ide_char Right_a.arr_char right_def Right_a.ide_def Right_a.null_char
+ using assms Right_a.ide_char\<^sub>S\<^sub>b\<^sub>C Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def Right_a.ide_def Right_a.null_char
by metis
obtain \<psi> where \<psi>: "\<psi> \<in> Right_a.hom f g \<and> \<psi> \<star> a = \<phi>"
using assms 1 2 3 R.is_full [of g f \<phi>] H\<^sub>R_def by auto
have "Right_a.iso \<psi>"
using \<psi> 1 H\<^sub>R_def R.reflects_iso by auto
hence "iso \<psi> \<and> \<guillemotleft>\<psi> : f \<Rightarrow> g\<guillemotright>"
- using \<psi> Right_a.iso_char Right_a.in_hom_char by auto
+ using \<psi> Right_a.iso_char Right_a.in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis by auto
qed
text \<open>
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.
\<close>
lemma source_iso_src:
assumes "arr \<mu>" and "a \<in> sources \<mu>"
shows "a \<cong> src \<mu>"
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 \<phi> where \<phi>: "iso \<phi> \<and> \<guillemotleft>\<phi> : a \<star> a \<Rightarrow> a\<guillemotright>"
using assms weak_unit_def by blast
have "a \<star> src a \<cong> src a \<star> src a"
proof -
have "src a \<cong> src a \<star> src a"
using 0 obj_is_weak_unit weak_unit_def isomorphic_symmetric by auto
moreover have "a \<star> src a \<cong> src a"
proof -
have "a \<star> a \<star> src a \<cong> a \<star> src a"
proof -
have "iso (\<phi> \<star> src a) \<and> \<guillemotleft>\<phi> \<star> src a : (a \<star> a) \<star> src a \<Rightarrow> a \<star> src a\<guillemotright>"
using 0 1 \<phi> ide_in_hom(2) by auto
moreover have "iso \<a>\<^sup>-\<^sup>1[a, a, src a] \<and>
\<guillemotleft>\<a>\<^sup>-\<^sup>1[a, a, src a] : a \<star> a \<star> src a \<Rightarrow> (a \<star> a) \<star> src a\<guillemotright>"
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 \<cong> 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 \<mu>" and "b \<in> targets \<mu>"
shows "b \<cong> trg \<mu>"
proof -
have 0: "ide b"
using assms weak_unit_def by force
have 1: "trg \<mu> = src b"
using assms seq_if_composable by auto
obtain \<phi> where \<phi>: "iso \<phi> \<and> \<guillemotleft>\<phi> : b \<star> b \<Rightarrow> b\<guillemotright>"
using assms weak_unit_def by blast
have "trg b \<star> b \<cong> trg b \<star> trg b"
proof -
have "trg b \<cong> trg b \<star> trg b"
using 0 obj_is_weak_unit weak_unit_def isomorphic_symmetric by auto
moreover have "trg b \<star> b \<cong> trg b"
proof -
have "(trg b \<star> b) \<star> b \<cong> trg b \<star> b"
proof -
have "iso (trg b \<star> \<phi>) \<and> \<guillemotleft>trg b \<star> \<phi> : trg b \<star> b \<star> b \<Rightarrow> trg b \<star> b\<guillemotright>"
using assms 0 1 \<phi> ide_in_hom(2) targetsD(1) weak_unit_self_composable
apply (intro conjI hcomp_in_vhom) by auto
moreover have "iso \<a>[trg b, b, b] \<and>
\<guillemotleft>\<a>[trg b, b, b] : (trg b \<star> b) \<star> b \<Rightarrow> trg b \<star> b \<star> b\<guillemotright>"
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 \<cong> 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 \<open>
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.
\<close>
lemma connected_iff_seq:
assumes "arr \<mu>" and "arr \<nu>"
shows "sources \<nu> \<inter> targets \<mu> \<noteq> {} \<longleftrightarrow> src \<nu> = trg \<mu>"
proof
show "src \<nu> = trg \<mu> \<Longrightarrow> sources \<nu> \<inter> targets \<mu> \<noteq> {}"
using assms src_in_sources [of \<nu>] trg_in_targets [of \<mu>] by auto
show "sources \<nu> \<inter> targets \<mu> \<noteq> {} \<Longrightarrow> src \<nu> = trg \<mu>"
proof -
assume 1: "sources \<nu> \<inter> targets \<mu> \<noteq> {}"
obtain a where a: "a \<in> sources \<nu> \<inter> targets \<mu>"
using assms 1 by blast
have \<mu>: "arr \<mu>"
using a composable_implies_arr by auto
have \<nu>: "arr \<nu>"
using a composable_implies_arr by auto
have 2: "\<And>a'. a' \<in> sources \<nu> \<Longrightarrow> src a' = src a \<and> trg a' = trg a"
by (metis IntE a seq_if_composable sourcesD(2-3) weak_unit_self_composable(3))
have "src \<nu> = src (src \<nu>)" using \<nu> by simp
also have "... = src (trg \<mu>)"
using \<nu> 2 [of "src \<nu>"] src_in_sources a weak_unit_self_composable seq_if_composable
by auto
also have "... = trg (trg \<mu>)" using \<mu> by simp
also have "... = trg \<mu>" using \<mu> by simp
finally show "src \<nu> = trg \<mu>" by blast
qed
qed
lemma is_associative_weak_composition:
shows "associative_weak_composition V H \<a>"
proof -
have 1: "\<And>\<nu> \<mu>. \<nu> \<star> \<mu> \<noteq> null \<Longrightarrow> src \<nu> = trg \<mu>"
- using H.is_extensional VV.arr_char by force
+ using H.is_extensional VV.arr_char\<^sub>S\<^sub>b\<^sub>C by force
show "associative_weak_composition V H \<a>"
proof
show "\<And>f g h. ide f \<Longrightarrow> ide g \<Longrightarrow> ide h \<Longrightarrow> f \<star> g \<noteq> null \<Longrightarrow> g \<star> h \<noteq> null \<Longrightarrow>
\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
using 1 by auto
show "\<And>f g h. ide f \<Longrightarrow> ide g \<Longrightarrow> ide h \<Longrightarrow> f \<star> g \<noteq> null \<Longrightarrow> g \<star> h \<noteq> null \<Longrightarrow>
iso \<a>[f, g, h]"
using 1 iso_assoc by presburger
show "\<And>\<tau> \<mu> \<nu>. \<tau> \<star> \<mu> \<noteq> null \<Longrightarrow> \<mu> \<star> \<nu> \<noteq> null \<Longrightarrow>
\<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>) = (\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>]"
using 1 assoc_naturality hseq_char hseq_char' by metis
show "\<And>f g h k. ide f \<Longrightarrow> ide g \<Longrightarrow> ide h \<Longrightarrow> ide k \<Longrightarrow>
sources f \<inter> targets g \<noteq> {} \<Longrightarrow>
sources g \<inter> targets h \<noteq> {} \<Longrightarrow>
sources h \<inter> targets k \<noteq> {} \<Longrightarrow>
(f \<star> \<a>[g, h, k]) \<cdot> \<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k) =
\<a>[f, g, h \<star> k] \<cdot> \<a>[f \<star> g, h, k]"
using 1 connected_iff_seq pentagon ideD(1) by auto
qed
qed
interpretation associative_weak_composition V H \<a>
using is_associative_weak_composition by auto
theorem is_prebicategory:
shows "prebicategory V H \<a>"
using src_in_sources trg_in_targets by (unfold_locales, auto)
interpretation prebicategory V H \<a>
using is_prebicategory by auto
corollary is_prebicategory_with_homs:
shows "prebicategory_with_homs V H \<a> src trg"
..
interpretation prebicategory_with_homs V H \<a> src trg
using is_prebicategory_with_homs by auto
text \<open>
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).
\<close>
lemma weak_unit_char:
shows "weak_unit a \<longleftrightarrow> a \<cong> src a"
and "weak_unit a \<longleftrightarrow> a \<cong> trg a"
proof -
show "weak_unit a \<longleftrightarrow> a \<cong> 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 \<longleftrightarrow> a \<cong> trg a"
using isomorphism_respects_weak_units isomorphic_symmetric
by (metis \<open>weak_unit a = isomorphic a (src a)\<close> 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 \<open>
Every arrow with respect to horizontal composition is also an arrow with respect
to vertical composition. The converse is not necessarily true.
\<close>
lemma harr_is_varr:
assumes "H.arr \<mu>"
shows "arr \<mu>"
using H.arr_def H.codomains_def H.domains_def assms composable_implies_arr(1)
composable_implies_arr(2)
by auto
text \<open>
An identity for horizontal composition is also an identity for vertical composition.
\<close>
lemma horizontal_identity_is_ide:
assumes "H.ide \<mu>"
shows "ide \<mu>"
proof -
have \<mu>: "arr \<mu>"
using assms H.ide_def composable_implies_arr(2) by auto
hence 1: "\<mu> \<star> dom \<mu> \<noteq> null"
using assms hom_connected H.ide_def by auto
have "\<mu> \<star> dom \<mu> = dom \<mu>"
using assms 1 H.ide_def by simp
moreover have "\<mu> \<star> dom \<mu> = \<mu>"
using assms 1 H.ide_def [of \<mu>] null_agreement
by (metis \<mu> cod_cod cod_dom hcomp_simps\<^sub>W\<^sub>C(3) ideD(2) ide_char' paste_1)
ultimately have "dom \<mu> = \<mu>"
by simp
thus ?thesis
using \<mu> by (metis ide_dom)
qed
text \<open>
Every identity for horizontal composition is a weak unit.
\<close>
lemma horizontal_identity_is_weak_unit:
assumes "H.ide \<mu>"
shows "weak_unit \<mu>"
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 \<open>
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.
\<close>
locale vertically_discrete_bicategory =
bicategory +
assumes vertically_discrete: "ide = arr"
begin
interpretation prebicategory_with_homs V H \<a> 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 \<star> a \<noteq> H.null"
using assms weak_unit_self_composable by simp
moreover have "\<And>f. f \<star> a \<noteq> H.null \<Longrightarrow> f \<star> a = f"
proof -
fix f
assume "f \<star> a \<noteq> H.null"
hence "f \<star> a \<cong> f"
using assms comp_ide_source composable_implies_arr(2) sourcesI vertically_discrete
by auto
thus "f \<star> a = f"
using vertically_discrete isomorphic_def by auto
qed
moreover have "\<And>f. a \<star> f \<noteq> H.null \<Longrightarrow> a \<star> f = f"
proof -
fix f
assume "a \<star> f \<noteq> H.null"
hence "a \<star> f \<cong> f"
using assms comp_target_ide composable_implies_arr(1) targetsI vertically_discrete
by auto
thus "a \<star> 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 \<mu> = H.domains \<mu>"
using weak_unit_is_horizontal_identity H.domains_def sources_def
horizontal_identity_is_weak_unit
by auto
lemma targets_eq_codomains:
shows "targets \<mu> = H.codomains \<mu>"
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 "\<And>g f. g \<star> f \<noteq> H.null \<Longrightarrow> H.seq g f"
using arr_agreement hcomp_simps\<^sub>W\<^sub>C(1) by auto
show "\<And>f. (H.domains f \<noteq> {}) = (H.codomains f \<noteq> {})"
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 \<Longrightarrow> H.seq (h \<star> g) f \<Longrightarrow> H.seq g f"
- using null_agreement arr_agreement H.not_arr_null preserves_arr VoV.arr_char
+ using null_agreement arr_agreement H.not_arr_null preserves_arr VoV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis hseq_char' match_1)
show "H.seq h (g \<star> f) \<Longrightarrow> H.seq g f \<Longrightarrow> H.seq h g"
- using null_agreement arr_agreement H.not_arr_null preserves_arr VoV.arr_char
+ using null_agreement arr_agreement H.not_arr_null preserves_arr VoV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis hseq_char' match_2)
show "H.seq g f \<Longrightarrow> H.seq h g \<Longrightarrow> H.seq (h \<star> g) f"
using arr_agreement match_3 hseq_char(1) by auto
show "H.seq g f \<Longrightarrow> H.seq h g \<Longrightarrow> (h \<star> g) \<star> f = h \<star> g \<star> f"
proof -
assume hg: "H.seq h g"
assume gf: "H.seq g f"
have "iso \<a>[h, g, f] \<and> \<guillemotleft>\<a>[h, g, f] : (h \<star> g) \<star> f \<Rightarrow> h \<star> g \<star> f\<guillemotright>"
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 \<open>
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 \<open>\<iota>\<close>, 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.
\<close>
context bicategory
begin
interpretation prebicategory V H \<a>
using is_prebicategory by auto
definition \<i>'
where "\<i>' a \<equiv> SOME \<phi>. iso \<phi> \<and> \<phi> \<in> hom (a \<star> a) a \<and> (obj a \<longrightarrow> \<phi> = \<i>[a])"
lemma \<i>'_extends_\<i>:
assumes "weak_unit a"
shows "iso (\<i>' a)" and "\<guillemotleft>\<i>' a : a \<star> a \<Rightarrow> a\<guillemotright>" and "obj a \<Longrightarrow> \<i>' a = \<i>[a]"
proof -
let ?P = "\<lambda>a \<phi>. iso \<phi> \<and> \<guillemotleft>\<phi> : a \<star> a \<Rightarrow> a\<guillemotright> \<and> (obj a \<longrightarrow> \<phi> = \<i>[a])"
have "\<exists>\<phi>. ?P a \<phi>"
by (metis assms iso_some_unit(1) iso_some_unit(2) iso_unit unit_in_vhom)
hence 1: "?P a (\<i>' a)"
using \<i>'_def someI_ex [of "?P a"] by simp
show "iso (\<i>' a)" using 1 by simp
show "\<guillemotleft>\<i>' a : a \<star> a \<Rightarrow> a\<guillemotright>" using 1 by simp
show "obj a \<Longrightarrow> \<i>' a = \<i>[a]" using 1 by simp
qed
proposition extends_to_prebicategory_with_units:
shows "prebicategory_with_units V H \<a> \<i>'"
using \<i>'_extends_\<i> by unfold_locales auto
interpretation PB: prebicategory_with_units V H \<a> \<i>'
using extends_to_prebicategory_with_units by auto
interpretation PB: prebicategory_with_homs V H \<a> src trg
using is_prebicategory_with_homs by auto
interpretation PB: prebicategory_with_homs_and_units V H \<a> \<i>' src trg ..
proposition extends_to_prebicategory_with_homs_and_units:
shows "prebicategory_with_homs_and_units V H \<a> \<i>' src trg"
..
definition lunit ("\<l>[_]")
where "\<l>[a] \<equiv> PB.lunit a"
definition runit ("\<r>[_]")
where "\<r>[a] \<equiv> PB.runit a"
abbreviation lunit' ("\<l>\<^sup>-\<^sup>1[_]")
where "\<l>\<^sup>-\<^sup>1[a] \<equiv> inv \<l>[a]"
abbreviation runit' ("\<r>\<^sup>-\<^sup>1[_]")
where "\<r>\<^sup>-\<^sup>1[a] \<equiv> inv \<r>[a]"
text \<open>
\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 \<open>\<i>'\<close>,
rather than the given \<open>\<i>\<close>. We want ``native versions'' for the present context.
\<close>
lemma lunit_char:
assumes "ide f"
shows "\<guillemotleft>\<l>[f] : L f \<Rightarrow> f\<guillemotright>" and "L \<l>[f] = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow> f\<guillemotright> \<and> L \<mu> = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^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 "\<guillemotleft>\<l>[f] : L f \<Rightarrow> f\<guillemotright>"
unfolding lunit_def
using assms PB.lunit_char by simp
show "L \<l>[f] = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]"
unfolding lunit_def
using assms 1 PB.lunit_char obj_is_weak_unit \<i>'_extends_\<i> by simp
let ?P = "\<lambda>\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow> f\<guillemotright> \<and> L \<mu> = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]"
have "?P = (\<lambda>\<mu>. \<guillemotleft>\<mu> : trg f \<star> f \<Rightarrow> f\<guillemotright> \<and>
trg f \<star> \<mu> = (\<i>' (trg f) \<star> f) \<cdot> inv \<a>[trg f, trg f, f])"
proof -
have "\<And>\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow> f\<guillemotright> \<longleftrightarrow> \<guillemotleft>\<mu> : trg f \<star> f \<Rightarrow> f\<guillemotright>"
using assms by simp
moreover have "\<And>\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow> f\<guillemotright> \<Longrightarrow>
L \<mu> = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, trg f, f] \<longleftrightarrow>
trg f \<star> \<mu> = (\<i>' (trg f) \<star> f) \<cdot> inv \<a>[trg f, trg f, f]"
using calculation obj_is_weak_unit \<i>'_extends_\<i> by auto
ultimately show ?thesis by blast
qed
thus "\<exists>!\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow> f\<guillemotright> \<and> L \<mu> = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^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 "\<guillemotleft>\<l>[f] : src f \<rightarrow> trg f\<guillemotright>" and "\<guillemotleft>\<l>[f] : trg f \<star> f \<Rightarrow> f\<guillemotright>"
proof -
show "\<guillemotleft>\<l>[f] : trg f \<star> f \<Rightarrow> f\<guillemotright>"
using assms lunit_char by auto
thus "\<guillemotleft>\<l>[f] : src f \<rightarrow> trg f\<guillemotright>"
by (metis arrI in_hhomI vconn_implies_hpar(1-4))
qed
lemma lunit_in_vhom [simp]:
assumes "ide f" and "trg f = b"
shows "\<guillemotleft>\<l>[f] : b \<star> f \<Rightarrow> f\<guillemotright>"
using assms by auto
lemma lunit_simps [simp]:
assumes "ide f"
shows "arr \<l>[f]" and "src \<l>[f] = src f" and "trg \<l>[f] = trg f"
and "dom \<l>[f] = trg f \<star> f" and "cod \<l>[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 "\<guillemotleft>\<r>[f] : R f \<Rightarrow> f\<guillemotright>" and "R \<r>[f] = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow> f\<guillemotright> \<and> R \<mu> = (f \<star> \<i>[src f]) \<cdot> \<a>[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 "\<guillemotleft>\<r>[f] : R f \<Rightarrow> f\<guillemotright>"
unfolding runit_def
using assms PB.runit_char by simp
show "R \<r>[f] = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
unfolding runit_def
using assms 1 PB.runit_char obj_is_weak_unit \<i>'_extends_\<i> by simp
let ?P = "\<lambda>\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow> f\<guillemotright> \<and> R \<mu> = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
have "?P = (\<lambda>\<mu>. \<guillemotleft>\<mu> : f \<star> src f \<Rightarrow> f\<guillemotright> \<and>
\<mu> \<star> src f = (f \<star> \<i>' (src f)) \<cdot> \<a>[f, src f, src f])"
proof -
have "\<And>\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow> f\<guillemotright> \<longleftrightarrow> \<guillemotleft>\<mu> : f \<star> src f \<Rightarrow> f\<guillemotright>"
using assms by simp
moreover have "\<And>\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow> f\<guillemotright> \<Longrightarrow>
R \<mu> = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f] \<longleftrightarrow>
\<mu> \<star> src f = (f \<star> \<i>' (src f)) \<cdot> \<a>[f, src f, src f]"
using calculation obj_is_weak_unit \<i>'_extends_\<i> by auto
ultimately show ?thesis by blast
qed
thus "\<exists>!\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow> f\<guillemotright> \<and> R \<mu> = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
using assms PB.runit_char by simp
qed
lemma runit_in_hom [intro]:
assumes "ide f"
shows "\<guillemotleft>\<r>[f] : src f \<rightarrow> trg f\<guillemotright>" and "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>"
proof -
show "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>"
using assms runit_char by auto
thus "\<guillemotleft>\<r>[f] : src f \<rightarrow> trg f\<guillemotright>"
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 "\<guillemotleft>\<r>[f] : f \<star> a \<Rightarrow> f\<guillemotright>"
using assms by auto
lemma runit_simps [simp]:
assumes "ide f"
shows "arr \<r>[f]" and "src \<r>[f] = src f" and "trg \<r>[f] = trg f"
and "dom \<r>[f] = f \<star> src f" and "cod \<r>[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 "\<guillemotleft>\<mu> : trg f \<star> f \<Rightarrow> f\<guillemotright>"
and "trg f \<star> \<mu> = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]"
shows "\<mu> = \<l>[f]"
unfolding lunit_def
using assms PB.lunit_eqI \<i>'_extends_\<i> trg.preserves_ide obj_is_weak_unit by simp
lemma runit_eqI:
assumes "ide f" and "\<guillemotleft>\<mu> : f \<star> src f \<Rightarrow> f\<guillemotright>"
and "\<mu> \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
shows "\<mu> = \<r>[f]"
unfolding runit_def
using assms PB.runit_eqI \<i>'_extends_\<i> src.preserves_ide obj_is_weak_unit by simp
lemma lunit_naturality:
assumes "arr \<mu>"
shows "\<mu> \<cdot> \<l>[dom \<mu>] = \<l>[cod \<mu>] \<cdot> (trg \<mu> \<star> \<mu>)"
unfolding lunit_def
using assms PB.lunit_naturality by auto
lemma runit_naturality:
assumes "arr \<mu>"
shows "\<mu> \<cdot> \<r>[dom \<mu>] = \<r>[cod \<mu>] \<cdot> (\<mu> \<star> src \<mu>)"
unfolding runit_def
using assms PB.runit_naturality by auto
lemma iso_lunit [simp]:
assumes "ide f"
shows "iso \<l>[f]"
unfolding lunit_def
using assms PB.iso_lunit by blast
lemma iso_runit [simp]:
assumes "ide f"
shows "iso \<r>[f]"
unfolding runit_def
using assms PB.iso_runit by blast
lemma iso_lunit' [simp]:
assumes "ide f"
shows "iso \<l>\<^sup>-\<^sup>1[f]"
using assms iso_lunit by blast
lemma iso_runit' [simp]:
assumes "ide f"
shows "iso \<r>\<^sup>-\<^sup>1[f]"
using assms iso_runit by blast
lemma lunit'_in_hom [intro]:
assumes "ide f"
shows "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] : src f \<rightarrow> trg f\<guillemotright>" and "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] : f \<Rightarrow> trg f \<star> f\<guillemotright>"
proof -
show "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] : f \<Rightarrow> trg f \<star> f\<guillemotright>"
using assms lunit_char iso_lunit by simp
thus "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] : src f \<rightarrow> trg f\<guillemotright>"
using assms src_dom trg_dom by simp
qed
lemma lunit'_in_vhom [simp]:
assumes "ide f" and "trg f = b"
shows "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] : f \<Rightarrow> b \<star> f\<guillemotright>"
using assms by auto
lemma lunit'_simps [simp]:
assumes "ide f"
shows "arr \<l>\<^sup>-\<^sup>1[f]" and "src \<l>\<^sup>-\<^sup>1[f] = src f" and "trg \<l>\<^sup>-\<^sup>1[f] = trg f"
and "dom \<l>\<^sup>-\<^sup>1[f] = f" and "cod \<l>\<^sup>-\<^sup>1[f] = trg f \<star> f"
using assms lunit'_in_hom by auto
lemma runit'_in_hom [intro]:
assumes "ide f"
shows "\<guillemotleft>\<r>\<^sup>-\<^sup>1[f] : src f \<rightarrow> trg f\<guillemotright>" and "\<guillemotleft>\<r>\<^sup>-\<^sup>1[f] : f \<Rightarrow> f \<star> src f\<guillemotright>"
proof -
show "\<guillemotleft>\<r>\<^sup>-\<^sup>1[f] : f \<Rightarrow> f \<star> src f\<guillemotright>"
using assms runit_char iso_runit by simp
thus "\<guillemotleft>\<r>\<^sup>-\<^sup>1[f] : src f \<rightarrow> trg f\<guillemotright>"
using src_dom trg_dom
by (simp add: assms)
qed
lemma runit'_in_vhom [simp]:
assumes "ide f" and "src f = a"
shows "\<guillemotleft>\<r>\<^sup>-\<^sup>1[f] : f \<Rightarrow> f \<star> a\<guillemotright>"
using assms by auto
lemma runit'_simps [simp]:
assumes "ide f"
shows "arr \<r>\<^sup>-\<^sup>1[f]" and "src \<r>\<^sup>-\<^sup>1[f] = src f" and "trg \<r>\<^sup>-\<^sup>1[f] = trg f"
and "dom \<r>\<^sup>-\<^sup>1[f] = f" and "cod \<r>\<^sup>-\<^sup>1[f] = f \<star> src f"
using assms runit'_in_hom by auto
interpretation L: endofunctor V L ..
interpretation \<ll>: transformation_by_components V V L map lunit
using lunit_in_hom lunit_naturality by unfold_locales auto
interpretation \<ll>: natural_isomorphism V V L map \<ll>.map
using iso_lunit by (unfold_locales, auto)
lemma natural_isomorphism_\<ll>:
shows "natural_isomorphism V V L map \<ll>.map"
..
abbreviation \<ll>
where "\<ll> \<equiv> \<ll>.map"
lemma \<ll>_ide_simp:
assumes "ide f"
shows "\<ll> f = \<l>[f]"
using assms by simp
interpretation L: equivalence_functor V V L
using L.isomorphic_to_identity_is_equivalence \<ll>.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>[L f] = L \<l>[f]"
unfolding lunit_def
using assms PB.lunit_commutes_with_L by blast
interpretation R: endofunctor V R ..
interpretation \<rr>: transformation_by_components V V R map runit
using runit_in_hom runit_naturality by unfold_locales auto
interpretation \<rr>: natural_isomorphism V V R map \<rr>.map
using iso_runit by (unfold_locales, auto)
lemma natural_isomorphism_\<rr>:
shows "natural_isomorphism V V R map \<rr>.map"
..
abbreviation \<rr>
where "\<rr> \<equiv> \<rr>.map"
lemma \<rr>_ide_simp:
assumes "ide f"
shows "\<rr> f = \<r>[f]"
using assms by simp
interpretation R: equivalence_functor V V R
using R.isomorphic_to_identity_is_equivalence \<rr>.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>[R f] = R \<r>[f]"
unfolding runit_def
using assms PB.runit_commutes_with_R by blast
lemma lunit'_naturality:
assumes "arr \<mu>"
shows "(trg \<mu> \<star> \<mu>) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu>] = \<l>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu>"
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 \<mu>"
shows "(\<mu> \<star> src \<mu>) \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>] = \<r>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu>"
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 \<star> src f \<cong> f"
using assms runit'_in_hom iso_runit' isomorphic_def isomorphic_symmetric by blast
lemma isomorphic_unit_left:
assumes "ide f"
shows "trg f \<star> f \<cong> f"
using assms lunit'_in_hom iso_lunit' isomorphic_def isomorphic_symmetric by blast
end
subsection "Further Properties of Bicategories"
text \<open>
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.
\<close>
context bicategory
begin
lemma unit_simps [simp]:
assumes "obj a"
shows "arr \<i>[a]" and "src \<i>[a] = a" and "trg \<i>[a] = a"
and "dom \<i>[a] = a \<star> a" and "cod \<i>[a] = a"
using assms unit_in_hom by blast+
lemma triangle:
assumes "ide f" and "ide g" and "src g = trg f"
shows "(g \<star> \<l>[f]) \<cdot> \<a>[g, src g, f] = \<r>[g] \<star> f"
proof -
let ?b = "src g"
have *: "(g \<star> \<l>[?b \<star> f]) \<cdot> \<a>[g, ?b, ?b \<star> f] = \<r>[g] \<star> ?b \<star> f"
proof -
have 1: "((g \<star> \<l>[?b \<star> f]) \<cdot> \<a>[g, ?b, ?b \<star> f]) \<cdot> \<a>[g \<star> ?b, ?b, f]
= (\<r>[g] \<star> ?b \<star> f) \<cdot> \<a>[g \<star> ?b, ?b, f]"
proof -
have "((g \<star> \<l>[?b \<star> f]) \<cdot> \<a>[g, ?b, ?b \<star> f]) \<cdot> \<a>[g \<star> ?b, ?b, f]
= (g \<star> \<l>[?b \<star> f]) \<cdot> \<a>[g, ?b, ?b \<star> f] \<cdot> \<a>[g \<star> ?b, ?b, f]"
using HoVH_def HoHV_def comp_assoc by auto
also have
"... = (g \<star> \<l>[?b \<star> f]) \<cdot> (g \<star> \<a>[?b, ?b, f]) \<cdot> \<a>[g, ?b \<star> ?b, f] \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms pentagon by force
also have
"... = ((g \<star> \<l>[?b \<star> f]) \<cdot> (g \<star> \<a>[?b, ?b, f])) \<cdot> \<a>[g, ?b \<star> ?b, f] \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms assoc_in_hom HoVH_def HoHV_def comp_assoc by auto
also have
"... = ((g \<star> ?b \<star> \<l>[f]) \<cdot> (g \<star> \<a>[?b, ?b, f])) \<cdot> \<a>[g, ?b \<star> ?b, f] \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms lunit_commutes_with_L lunit_in_hom by force
also have "... = ((g \<star> (\<i>[?b] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[?b, ?b, f]) \<cdot> (g \<star> \<a>[?b, ?b, f]))
\<cdot> \<a>[g, ?b \<star> ?b, f] \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms lunit_char(2) by force
also have "... = (g \<star> ((\<i>[?b] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[?b, ?b, f]) \<cdot> \<a>[?b, ?b, f])
\<cdot> \<a>[g, ?b \<star> ?b, f] \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms interchange [of g g "(\<i>[?b] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[?b, ?b, f]" "\<a>[?b, ?b, f]"]
by auto
also have "... = ((g \<star> \<i>[?b] \<star> f) \<cdot> \<a>[g, ?b \<star> ?b, f]) \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms comp_arr_dom comp_assoc_assoc' comp_assoc by auto
also have "... = (\<a>[g, ?b, f] \<cdot> ((g \<star> \<i>[?b]) \<star> f)) \<cdot> (\<a>[g, ?b, ?b] \<star> f)"
using assms assoc_naturality [of g "\<i>[?b]" f] by simp
also have "... = \<a>[g, ?b, f] \<cdot> ((g \<star> \<i>[?b]) \<cdot> \<a>[g, ?b, ?b] \<star> f)"
using assms interchange [of "g \<star> \<i>[?b]" "\<a>[g, ?b, ?b]" f f] comp_assoc by simp
also have "... = \<a>[g, ?b, f] \<cdot> ((\<r>[g] \<star> ?b) \<star> f)"
using assms runit_char(2) by force
also have "... = (\<r>[g] \<star> ?b \<star> f) \<cdot> \<a>[g \<star> ?b, ?b, f]"
using assms assoc_naturality [of "\<r>[g]" ?b f] by auto
finally show ?thesis by blast
qed
show "(g \<star> \<l>[?b \<star> f]) \<cdot> \<a>[g, ?b, ?b \<star> f] = \<r>[g] \<star> ?b \<star> f"
proof -
have "epi \<a>[g \<star> ?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 \<star> \<l>[f]) \<cdot> \<a>[g, ?b, f] = ((g \<star> \<l>[f]) \<cdot> (g \<star> \<l>[?b \<star> f]) \<cdot> (g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f])) \<cdot>
(g \<star> ?b \<star> \<l>[f]) \<cdot> \<a>[g, ?b, ?b \<star> f] \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
proof -
have "\<a>[g, ?b, f] = (g \<star> ?b \<star> \<l>[f]) \<cdot> \<a>[g, ?b, ?b \<star> f] \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
proof -
have "\<a>[g, ?b, f] = (g \<star> ?b \<star> f) \<cdot> \<a>[g, ?b, f]"
using assms comp_cod_arr by simp
have "\<a>[g, ?b, f] = ((g \<star> ?b \<star> \<l>[f]) \<cdot> (g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f])) \<cdot> \<a>[g, ?b, f]"
using assms comp_cod_arr comp_arr_inv' whisker_left [of g]
whisker_left [of ?b "\<l>[f]" "\<l>\<^sup>-\<^sup>1[f]"]
by simp
also have "... = (g \<star> ?b \<star> \<l>[f]) \<cdot> \<a>[g, ?b, ?b \<star> f] \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
using assms iso_lunit assoc_naturality [of g ?b "\<l>\<^sup>-\<^sup>1[f]"] comp_assoc by force
finally show ?thesis by blast
qed
moreover have "g \<star> \<l>[f] = (g \<star> \<l>[f]) \<cdot> (g \<star> \<l>[?b \<star> f]) \<cdot> (g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f])"
proof -
have "(g \<star> \<l>[?b \<star> f]) \<cdot> (g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f]) = g \<star> ?b \<star> f"
proof -
have "(g \<star> \<l>[?b \<star> f]) \<cdot> (g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f]) = (g \<star> ?b \<star> \<l>[f]) \<cdot> (g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f])"
using assms lunit_in_hom lunit_commutes_with_L by simp
also have "... = g \<star> ?b \<star> f"
using assms comp_arr_inv' whisker_left [of g] whisker_left [of ?b "\<l>[f]" "\<l>\<^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 \<star> \<l>[f]) \<cdot> (g \<star> \<l>[?b \<star> f]) \<cdot> ((g \<star> ?b \<star> \<l>\<^sup>-\<^sup>1[f]) \<cdot> (g \<star> ?b \<star> \<l>[f])) \<cdot>
\<a>[g, ?b, ?b \<star> f] \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
using comp_assoc by simp
also have "... = (g \<star> \<l>[f]) \<cdot> (g \<star> \<l>[?b \<star> f]) \<cdot> ((g \<star> ?b \<star> (?b \<star> f)) \<cdot>
\<a>[g, ?b, ?b \<star> f]) \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
using assms iso_lunit comp_inv_arr' interchange [of g g "?b \<star> \<l>\<^sup>-\<^sup>1[f]" "?b \<star> \<l>[f]"]
interchange [of ?b ?b "\<l>\<^sup>-\<^sup>1[f]" "\<l>[f]"] comp_assoc
by auto
also have "... = (g \<star> \<l>[f]) \<cdot> ((g \<star> \<l>[?b \<star> f]) \<cdot> \<a>[g, ?b, ?b \<star> f]) \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
using assms comp_cod_arr comp_assoc by auto
also have "... = \<r>[g] \<star> f"
proof -
have "\<r>[g] \<star> f = (g \<star> \<l>[f]) \<cdot> (\<r>[g] \<star> ?b \<star> f) \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])"
proof -
have "(g \<star> \<l>[f]) \<cdot> (\<r>[g] \<star> ?b \<star> f) \<cdot> ((g \<star> ?b) \<star> \<l>\<^sup>-\<^sup>1[f])
= (g \<star> \<l>[f]) \<cdot> (\<r>[g] \<cdot> (g \<star> ?b) \<star> (?b \<star> f) \<cdot> \<l>\<^sup>-\<^sup>1[f])"
using assms iso_lunit interchange [of "\<r>[g]" "g \<star> ?b" "?b \<star> f" "\<l>\<^sup>-\<^sup>1[f]"]
by force
also have "... = (g \<star> \<l>[f]) \<cdot> (\<r>[g] \<star> \<l>\<^sup>-\<^sup>1[f])"
using assms comp_arr_dom comp_cod_arr by simp
also have "... = \<r>[g] \<star> \<l>[f] \<cdot> \<l>\<^sup>-\<^sup>1[f]"
using assms interchange [of g "\<r>[g]" "\<l>[f]" "\<l>\<^sup>-\<^sup>1[f]"] comp_cod_arr
by simp
also have "... = \<r>[g] \<star> 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 \<star> \<l>[g \<star> h]) \<cdot> (f \<star> \<a>[trg g, g, h]) = f \<star> \<l>[g] \<star> h"
proof -
have "((f \<star> \<l>[g \<star> h]) \<cdot> (f \<star> \<a>[trg g, g, h])) \<cdot> \<a>[f, trg g \<star> g, h] \<cdot> (\<a>[f, trg g, g] \<star> h) =
(f \<star> (\<l>[g] \<star> h)) \<cdot> \<a>[f, trg g \<star> g, h] \<cdot> (\<a>[f, trg g, g] \<star> h)"
proof -
have "((f \<star> \<l>[g \<star> h]) \<cdot> (f \<star> \<a>[trg g, g, h])) \<cdot> (\<a>[f, trg g \<star> g, h] \<cdot> (\<a>[f, trg g, g] \<star> h)) =
((f \<star> \<l>[g \<star> h]) \<cdot> \<a>[f, trg g, g \<star> h]) \<cdot> \<a>[f \<star> trg g, g, h]"
using assms pentagon comp_assoc by simp
also have "... = (\<r>[f] \<star> (g \<star> h)) \<cdot> \<a>[f \<star> trg g, g, h]"
using assms triangle [of "g \<star> h" f] by auto
also have "... = \<a>[f, g, h] \<cdot> ((\<r>[f] \<star> g) \<star> h)"
using assms assoc_naturality [of "\<r>[f]" g h] by simp
also have "... = (\<a>[f, g, h] \<cdot> ((f \<star> \<l>[g]) \<star> h)) \<cdot> (\<a>[f, trg g, g] \<star> h)"
using assms triangle interchange [of "f \<star> \<l>[g]" "\<a>[f, trg g, g]" h h] comp_assoc
by auto
also have "... = (f \<star> (\<l>[g] \<star> h)) \<cdot> (\<a>[f, trg g \<star> g, h] \<cdot> (\<a>[f, trg g, g] \<star> h))"
using assms assoc_naturality [of f "\<l>[g]" h] comp_assoc by simp
finally show ?thesis by blast
qed
moreover have "iso (\<a>[f, trg g \<star> g, h] \<cdot> (\<a>[f, trg g, g] \<star> h))"
using assms iso_assoc isos_compose by simp
ultimately show ?thesis
using assms iso_is_retraction retraction_is_epi
epiE [of "\<a>[f, trg g \<star> g, h] \<cdot> (\<a>[f, trg g, g] \<star> h)"
"(f \<star> \<l>[g \<star> h]) \<cdot> (f \<star> \<a>[trg g, g, h])" "f \<star> \<l>[g] \<star> h"]
by auto
qed
lemma lunit_hcomp:
assumes "ide f" and "ide g" and "src f = trg g"
shows "\<l>[f \<star> g] \<cdot> \<a>[trg f, f, g] = \<l>[f] \<star> g"
and "\<a>\<^sup>-\<^sup>1[trg f, f, g] \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> g] = \<l>\<^sup>-\<^sup>1[f] \<star> g"
and "\<l>[f \<star> g] = (\<l>[f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, f, g]"
and "\<l>\<^sup>-\<^sup>1[f \<star> g] = \<a>[trg f, f, g] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<star> g)"
proof -
show 1: "\<l>[f \<star> g] \<cdot> \<a>[trg f, f, g] = \<l>[f] \<star> g"
proof -
have "L (\<l>[f \<star> g] \<cdot> \<a>[trg f, f, g]) = L (\<l>[f] \<star> g)"
using assms interchange [of "trg f" "trg f" "\<l>[f \<star> g]" "\<a>[trg f, f, g]"] lunit_hcomp_gen
by fastforce
thus ?thesis
using assms L.is_faithful [of "\<l>[f \<star> g] \<cdot> \<a>[trg f, f, g]" "\<l>[f] \<star> g"] by force
qed
show "\<a>\<^sup>-\<^sup>1[trg f, f, g] \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> g] = \<l>\<^sup>-\<^sup>1[f] \<star> g"
proof -
have "\<a>\<^sup>-\<^sup>1[trg f, f, g] \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> g] = inv (\<l>[f \<star> g] \<cdot> \<a>[trg f, f, g])"
using assms by (simp add: inv_comp)
also have "... = inv (\<l>[f] \<star> g)"
using 1 by simp
also have "... = \<l>\<^sup>-\<^sup>1[f] \<star> g"
using assms by simp
finally show ?thesis by simp
qed
show 2: "\<l>[f \<star> g] = (\<l>[f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, f, g]"
using assms 1 invert_side_of_triangle(2) by auto
show "\<l>\<^sup>-\<^sup>1[f \<star> g] = \<a>[trg f, f, g] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<star> g)"
proof -
have "\<l>\<^sup>-\<^sup>1[f \<star> g] = inv ((\<l>[f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, f, g])"
using 2 by simp
also have "... = \<a>[trg f, f, g] \<cdot> inv (\<l>[f] \<star> g)"
using assms inv_comp by simp
also have "... = \<a>[trg f, f, g] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<star> 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 "\<r>[f \<star> g] \<star> h = ((f \<star> \<r>[g]) \<star> h) \<cdot> (\<a>[f, g, src g] \<star> h)"
proof -
have "\<r>[f \<star> g] \<star> h = ((f \<star> g) \<star> \<l>[h]) \<cdot> \<a>[f \<star> g, src g, h]"
using assms triangle by simp
also have "... = (\<a>\<^sup>-\<^sup>1[f, g, h] \<cdot> (f \<star> g \<star> \<l>[h]) \<cdot> \<a>[f, g, src g \<star> h]) \<cdot> \<a>[f \<star> g, src g, h]"
using assms assoc_naturality [of f g "\<l>[h]"] invert_side_of_triangle(1)
by simp
also have "... = \<a>\<^sup>-\<^sup>1[f, g, h] \<cdot> (f \<star> g \<star> \<l>[h]) \<cdot> \<a>[f, g, src g \<star> h] \<cdot> \<a>[f \<star> g, src g, h]"
using comp_assoc by simp
also have "... = (\<a>\<^sup>-\<^sup>1[f, g, h] \<cdot> (f \<star> (\<r>[g] \<star> h))) \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[g, src g, h]) \<cdot>
\<a>[f, g, src g \<star> h] \<cdot> \<a>[f \<star> g, src g, h]"
using assms interchange [of f f] triangle comp_assoc
invert_side_of_triangle(2) [of "\<r>[g] \<star> h" "g \<star> \<l>[h]" "\<a>[g, src g, h]"]
by simp
also have "... = ((f \<star> \<r>[g]) \<star> h) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> src g, h] \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[g, src g, h]) \<cdot>
\<a>[f, g, src g \<star> h] \<cdot> \<a>[f \<star> g, src g, h]"
using assms assoc'_naturality [of f "\<r>[g]" h] comp_assoc by simp
also have "... = ((f \<star> \<r>[g]) \<star> h) \<cdot> (\<a>[f, g, src g] \<star> h)"
using assms pentagon [of f g "src g" h] iso_assoc inv_hcomp
invert_side_of_triangle(1)
[of "\<a>[f, g, src g \<star> h] \<cdot> \<a>[f \<star> g, src g, h]" "f \<star> \<a>[g, src g, h]"
"\<a>[f, g \<star> src g, h] \<cdot> (\<a>[f, g, src g] \<star> h)"]
invert_side_of_triangle(1)
[of "(f \<star> \<a>\<^sup>-\<^sup>1[g, src g, h]) \<cdot> \<a>[f, g, src g \<star> h] \<cdot> \<a>[f \<star> g, src g, h]"
"\<a>[f, g \<star> src g, h]" "\<a>[f, g, src g] \<star> h"]
by auto
finally show ?thesis by blast
qed
lemma runit_hcomp:
assumes "ide f" and "ide g" and "src f = trg g"
shows "\<r>[f \<star> g] = (f \<star> \<r>[g]) \<cdot> \<a>[f, g, src g]"
and "\<r>\<^sup>-\<^sup>1[f \<star> g] = \<a>\<^sup>-\<^sup>1[f, g, src g] \<cdot> (f \<star> \<r>\<^sup>-\<^sup>1[g])"
and "\<r>[f \<star> g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, src g] = f \<star> \<r>[g]"
and "\<a>[f, g, src g] \<cdot> \<r>\<^sup>-\<^sup>1[f \<star> g] = f \<star> \<r>\<^sup>-\<^sup>1[g]"
proof -
show 1: "\<r>[f \<star> g] = (f \<star> \<r>[g]) \<cdot> \<a>[f, g, src g]"
using assms interchange [of "f \<star> \<r>[g]" "\<a>[f, g, src g]" "src g" "src g"]
runit_hcomp_gen [of f g "src g"]
R.is_faithful [of "(f \<star> \<r>[g]) \<cdot> (\<a>[f, g, src g])" "\<r>[f \<star> g]"]
by simp
show "\<r>\<^sup>-\<^sup>1[f \<star> g] = \<a>\<^sup>-\<^sup>1[f, g, src g] \<cdot> (f \<star> \<r>\<^sup>-\<^sup>1[g])"
using assms 1 inv_comp inv_hcomp by auto
show 2: "\<r>[f \<star> g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, src g] = f \<star> \<r>[g]"
using assms 1 comp_arr_dom comp_cod_arr comp_assoc hseqI' comp_assoc_assoc' by auto
show "\<a>[f, g, src g] \<cdot> \<r>\<^sup>-\<^sup>1[f \<star> g] = f \<star> \<r>\<^sup>-\<^sup>1[g]"
proof -
have "\<a>[f, g, src g] \<cdot> \<r>\<^sup>-\<^sup>1[f \<star> g] = inv (\<r>[f \<star> g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, src g])"
using assms inv_comp by simp
also have "... = inv (f \<star> \<r>[g])"
using 2 by simp
also have "... = f \<star> \<r>\<^sup>-\<^sup>1[g]"
using assms inv_hcomp [of f "\<r>[g]"] by simp
finally show ?thesis by simp
qed
qed
lemma unitor_coincidence:
assumes "obj a"
shows "\<l>[a] = \<i>[a]" and "\<r>[a] = \<i>[a]"
proof -
have "R \<l>[a] = R \<i>[a] \<and> R \<r>[a] = R \<i>[a]"
proof -
have "R \<l>[a] = (a \<star> \<l>[a]) \<cdot> \<a>[a, a, a]"
using assms lunit_hcomp [of a a] lunit_commutes_with_L [of a] by auto
moreover have "(a \<star> \<l>[a]) \<cdot> \<a>[a, a, a] = R \<r>[a]"
using assms triangle [of a a] by auto
moreover have "(a \<star> \<l>[a]) \<cdot> \<a>[a, a, a] = R \<i>[a]"
proof -
have "(a \<star> \<l>[a]) \<cdot> \<a>[a, a, a] = ((\<i>[a] \<star> a) \<cdot> \<a>\<^sup>-\<^sup>1[a, a, a]) \<cdot> \<a>[a, a, a]"
using assms lunit_char(2) by force
also have "... = R \<i>[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 \<l>[a] \<i>[a] \<and> par \<r>[a] \<i>[a]"
using assms by auto
ultimately have 1: "\<l>[a] = \<i>[a] \<and> \<r>[a] = \<i>[a]"
using R.is_faithful by blast
show "\<l>[a] = \<i>[a]" using 1 by auto
show "\<r>[a] = \<i>[a]" using 1 by auto
qed
lemma unit_triangle:
assumes "obj a"
shows "\<i>[a] \<star> a = (a \<star> \<i>[a]) \<cdot> \<a>[a, a, a]"
and "(\<i>[a] \<star> a) \<cdot> \<a>\<^sup>-\<^sup>1[a, a, a] = a \<star> \<i>[a]"
proof -
show 1: "\<i>[a] \<star> a = (a \<star> \<i>[a]) \<cdot> \<a>[a, a, a]"
using assms triangle [of a a] unitor_coincidence by auto
show "(\<i>[a] \<star> a) \<cdot> \<a>\<^sup>-\<^sup>1[a, a, a] = a \<star> \<i>[a]"
using assms 1 invert_side_of_triangle(2) [of "\<i>[a] \<star> a" "a \<star> \<i>[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 \<star> g) \<star> h \<cong> f \<star> g \<star> h"
using assms assoc_in_hom [of f g h] iso_assoc isomorphic_def by auto
lemma hcomp_arr_obj:
assumes "arr \<mu>" and "obj a" and "src \<mu> = a"
shows "\<mu> \<star> a = \<r>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu> \<cdot> \<r>[dom \<mu>]"
and "\<r>[cod \<mu>] \<cdot> (\<mu> \<star> a) \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>] = \<mu>"
proof -
show "\<mu> \<star> a = \<r>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu> \<cdot> \<r>[dom \<mu>]"
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 "\<r>[cod \<mu>] \<cdot> (\<mu> \<star> a) \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>] = \<mu>"
using assms iso_runit runit_naturality [of \<mu>] 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 \<mu>" and "obj b" and "b = trg \<mu>"
shows "b \<star> \<mu> = \<l>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu> \<cdot> \<l>[dom \<mu>]"
and "\<l>[cod \<mu>] \<cdot> (b \<star> \<mu>) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu>] = \<mu>"
proof -
show "b \<star> \<mu> = \<l>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu> \<cdot> \<l>[dom \<mu>]"
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 "\<l>[cod \<mu>] \<cdot> (b \<star> \<mu>) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu>] = \<mu>"
using assms iso_lunit lunit_naturality [of \<mu>] 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 \<tau>" and "arr \<mu>" and "arr \<nu>"
and "src \<tau> = trg \<mu>" and "src \<mu> = trg \<nu>"
shows "(\<tau> \<star> \<mu>) \<star> \<nu> = \<a>\<^sup>-\<^sup>1[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> (\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>]"
and "\<tau> \<star> \<mu> \<star> \<nu> = \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<tau>, dom \<mu>, dom \<nu>]"
proof -
show "(\<tau> \<star> \<mu>) \<star> \<nu> = \<a>\<^sup>-\<^sup>1[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> (\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>]"
proof -
have "(\<tau> \<star> \<mu>) \<star> \<nu> = (\<a>\<^sup>-\<^sup>1[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> \<a>[cod \<tau>, cod \<mu>, cod \<nu>]) \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>)"
using assms comp_assoc_assoc'(2) comp_cod_arr by simp
also have "... = \<a>\<^sup>-\<^sup>1[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>)"
using comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> (\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>]"
using assms assoc_naturality by simp
finally show ?thesis by simp
qed
show "\<tau> \<star> \<mu> \<star> \<nu> = \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<tau>, dom \<mu>, dom \<nu>]"
proof -
have "\<tau> \<star> \<mu> \<star> \<nu> = (\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>] \<cdot> \<a>\<^sup>-\<^sup>1[dom \<tau>, dom \<mu>, dom \<nu>]"
using assms comp_assoc_assoc'(1) comp_arr_dom by simp
also have "... = ((\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>]) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<tau>, dom \<mu>, dom \<nu>]"
using comp_assoc by simp
also have "... = (\<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>)) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<tau>, dom \<mu>, dom \<nu>]"
using assms assoc_naturality by simp
also have "... = \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<tau>, dom \<mu>, dom \<nu>]"
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 \<star> \<l>[g]) = (\<r>[f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, g]"
using assms(1-3) invert_side_of_triangle(2) triangle by force
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 "((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> k) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> h, k]) \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[g, h, k])
= \<a>\<^sup>-\<^sup>1[f \<star> g, h, k] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, h \<star> k]"
proof -
have "((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> k) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> h, k]) \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[g, h, k])
= inv ((f \<star> \<a>[g, h, k]) \<cdot> (\<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k)))"
proof -
have "inv ((f \<star> \<a>[g, h, k]) \<cdot> (\<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k))) =
inv (\<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k)) \<cdot> inv (f \<star> \<a>[g, h, k])"
using assms inv_comp [of "\<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k)" "f \<star> \<a>[g, h, k]"]
by force
also have "... = (inv (\<a>[f, g, h] \<star> k) \<cdot> inv \<a>[f, g \<star> h, k]) \<cdot> inv (f \<star> \<a>[g, h, k])"
using assms iso_assoc inv_comp by simp
also have "... = ((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> k) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> h, k]) \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[g, h, k])"
using assms inv_hcomp by simp
finally show ?thesis by simp
qed
also have "... = inv (\<a>[f, g, h \<star> k] \<cdot> \<a>[f \<star> g, h, k])"
using assms pentagon by simp
also have "... = \<a>\<^sup>-\<^sup>1[f \<star> g, h, k] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, h \<star> k]"
using assms inv_comp by simp
finally show ?thesis by auto
qed
end
text \<open>
The following convenience locale extends @{locale bicategory} by pre-interpreting
the various functors and natural transformations.
\<close>
locale extended_bicategory =
bicategory +
L: equivalence_functor V V L +
R: equivalence_functor V V R +
\<alpha>: natural_isomorphism VVV.comp V HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close> +
\<alpha>': inverse_transformation VVV.comp V HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close> +
\<ll>: natural_isomorphism V V L map \<ll> +
\<ll>': inverse_transformation V V L map \<ll> +
\<rr>: natural_isomorphism V V R map \<rr> +
\<rr>': inverse_transformation V V R map \<rr>
sublocale bicategory \<subseteq> extended_bicategory V H \<a> \<i> 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 \<alpha>': inverse_transformation VVV.comp V HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))\<close> ..
interpret \<ll>: natural_isomorphism V V L map \<ll> using natural_isomorphism_\<ll> by auto
interpret \<ll>': inverse_transformation V V L map \<ll> ..
interpret \<rr>: natural_isomorphism V V R map \<rr> using natural_isomorphism_\<rr> by auto
interpret \<rr>': inverse_transformation V V R map \<rr> ..
interpret extended_bicategory V H \<a> \<i> src trg ..
show "extended_bicategory V H \<a> \<i> 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,14655 +1,14655 @@
(* Title: BicategoryOfSpans
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Bicategories of Spans"
theory BicategoryOfSpans
imports Category3.ConcreteCategory IsomorphismClass CanonicalIsos EquivalenceOfBicategories
SpanBicategory Tabulation
begin
text \<open>
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 \<open>g \<star> f\<^sup>*\<close>, where f and g are maps'';
BS2: ``For every span of maps \<open>(f, g)\<close> there is a 2-cell \<open>\<rho>\<close> such that \<open>(f, \<rho>, g)\<close>
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 \<open>B\<close>
is biequivalent to the bicategory of spans in some category \<open>C\<close> with pullbacks,
then it satisfies BS1 -- BS3.
The other direction, which is harder, shows that a bicategory \<open>B\<close> 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 \<open>B\<close>.
\<close>
subsection "Definition"
text \<open>
We define a \emph{bicategory of spans} to be a bicategory that satisfies the conditions
\<open>BS1\<close> -- \<open>BS3\<close> stated informally above.
\<close>
locale bicategory_of_spans =
bicategory + chosen_right_adjoints +
assumes BS1: "ide r \<Longrightarrow> \<exists>f g. is_left_adjoint f \<and> is_left_adjoint g \<and> isomorphic r (g \<star> f\<^sup>*)"
and BS2: "\<lbrakk> is_left_adjoint f; is_left_adjoint g; src f = src g \<rbrakk>
\<Longrightarrow> \<exists>r \<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"
and BS3: "\<lbrakk> is_left_adjoint f; is_left_adjoint f'; \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright>; \<guillemotleft>\<mu>' : f \<Rightarrow> f'\<guillemotright> \<rbrakk>
\<Longrightarrow> iso \<mu> \<and> iso \<mu>' \<and> \<mu> = \<mu>'"
text \<open>
Using the already-established fact \<open>equivalence_pseudofunctor.reflects_tabulation\<close>
that tabulations are reflected by equivalence pseudofunctors, it is not difficult to prove
that the notion `bicategory of spans' respects equivalence of bicategories.
\<close>
lemma bicategory_of_spans_respects_equivalence:
assumes "equivalent_bicategories V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
and "bicategory_of_spans V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C"
shows "bicategory_of_spans V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
proof -
interpret C: bicategory_of_spans V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
using assms by simp
interpret C: chosen_right_adjoints V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C ..
interpret D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^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 \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
obtain F \<Phi> where F: "equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>"
using assms equivalent_bicategories_def by blast
interpret F: equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>
using F by simp
interpret E: equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C (* 17 sec *)
F \<Phi> 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 \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi> 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 \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^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 "\<cdot>\<^sub>C" 55)
write V\<^sub>D (infixr "\<cdot>\<^sub>D" 55)
write H\<^sub>C (infixr "\<star>\<^sub>C" 53)
write H\<^sub>D (infixr "\<star>\<^sub>D" 53)
write \<a>\<^sub>C ("\<a>\<^sub>C[_, _, _]")
write \<a>\<^sub>D ("\<a>\<^sub>D[_, _, _]")
write C.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
write C.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>C _\<guillemotright>")
write D.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>D _\<guillemotright>")
write D.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>D _\<guillemotright>")
write C.isomorphic (infix "\<cong>\<^sub>C" 50)
write D.isomorphic (infix "\<cong>\<^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 \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
proof
show "\<And>r'. D.ide r' \<Longrightarrow>
\<exists>f' g'. D.is_left_adjoint f' \<and> D.is_left_adjoint g' \<and> r' \<cong>\<^sub>D g' \<star>\<^sub>D (f')\<^sup>*\<^sup>D"
proof -
fix r'
assume r': "D.ide r'"
obtain f g
where fg: "C.is_left_adjoint f \<and> C.is_left_adjoint g \<and> F.right_map r' \<cong>\<^sub>C g \<star>\<^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 \<equiv> F.counit\<^sub>0 (src\<^sub>D r')"
define e_src where "e_src \<equiv> (F.counit\<^sub>0 (src\<^sub>D r'))\<^sup>~\<^sup>D"
have d_src: "\<guillemotleft>d_src : F.map\<^sub>0 (E.G.map\<^sub>0 (src\<^sub>D r')) \<rightarrow>\<^sub>D src\<^sub>D r'\<guillemotright> \<and>
D.equivalence_map d_src"
using d_src_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
have e_src: "\<guillemotleft>e_src : src\<^sub>D r' \<rightarrow>\<^sub>D F.map\<^sub>0 (E.G.map\<^sub>0 (src\<^sub>D r'))\<guillemotright> \<and>
D.equivalence_map e_src"
using e_src_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
obtain \<eta>_src \<epsilon>_src
where eq_src: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \<eta>_src \<epsilon>_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 \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \<eta>_src \<epsilon>_src
using eq_src by simp
define d_trg where "d_trg \<equiv> F.counit\<^sub>0 (trg\<^sub>D r')"
define e_trg where "e_trg \<equiv> (F.counit\<^sub>0 (trg\<^sub>D r'))\<^sup>~\<^sup>D"
have d_trg: "\<guillemotleft>d_trg : F.map\<^sub>0 (E.G.map\<^sub>0 (trg\<^sub>D r')) \<rightarrow>\<^sub>D trg\<^sub>D r'\<guillemotright> \<and>
D.equivalence_map d_trg"
using d_trg_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
have e_trg: "\<guillemotleft>e_trg : trg\<^sub>D r' \<rightarrow>\<^sub>D F.map\<^sub>0 (E.G.map\<^sub>0 (trg\<^sub>D r'))\<guillemotright> \<and>
D.equivalence_map e_trg"
using e_trg_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
obtain \<eta>_trg \<epsilon>_trg
where eq_trg: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_trg e_trg \<eta>_trg \<epsilon>_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 \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_trg e_trg \<eta>_trg \<epsilon>_trg
using eq_trg by simp
interpret eqs: two_equivalences_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
d_src e_src \<eta>_src \<epsilon>_src d_trg e_trg \<eta>_trg \<epsilon>_trg
..
interpret hom: subcategory V\<^sub>D \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : trg\<^sub>D d_src \<rightarrow>\<^sub>D trg\<^sub>D d_trg\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
interpret hom': subcategory V\<^sub>D \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : src\<^sub>D d_src \<rightarrow>\<^sub>D src\<^sub>D d_trg\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
interpret e: equivalence_of_categories hom.comp hom'.comp eqs.F eqs.G eqs.\<phi> eqs.\<psi>
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 \<star>\<^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 \<star>\<^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 \<star>\<^sub>D e_src)"
using 1 by auto
have f': "D.is_left_adjoint f' \<and> F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src \<cong>\<^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' \<cong>\<^sub>D d_trg \<star>\<^sub>D (e_trg \<star>\<^sub>D r' \<star>\<^sub>D d_src) \<star>\<^sub>D e_src"
using r' r'_in_hhom D.isomorphic_def eqs.\<psi>_in_hom eqs.\<psi>_components_are_iso
D.isomorphic_symmetric D.ide_char eq_src.antipar(2) eq_trg.antipar(2)
by metis
also have 1: "... \<cong>\<^sub>D d_trg \<star>\<^sub>D F (F.right_map r') \<star>\<^sub>D e_src"
proof -
have "e_trg \<star>\<^sub>D r' \<star>\<^sub>D d_src \<cong>\<^sub>D F (F.right_map r')"
proof -
have "D.in_hom (F.counit\<^sub>1 r')
(r' \<star>\<^sub>D d_src) (F.counit\<^sub>0 (trg\<^sub>D r') \<star>\<^sub>D F (F.right_map r'))"
unfolding d_src_def
using r' E.\<epsilon>.map\<^sub>1_in_hom(2) [of r'] by simp
hence "r' \<star>\<^sub>D d_src \<cong>\<^sub>D F.counit\<^sub>0 (trg\<^sub>D r') \<star>\<^sub>D F (F.right_map r')"
using r' D.isomorphic_def E.\<epsilon>.iso_map\<^sub>1_ide by auto
thus ?thesis
using r' e_trg_def E.\<epsilon>.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: "... \<cong>\<^sub>D d_trg \<star>\<^sub>D (F g \<star>\<^sub>D F f\<^sup>*\<^sup>C) \<star>\<^sub>D e_src"
proof -
have "F (F.right_map r') \<cong>\<^sub>D F g \<star>\<^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: "... \<cong>\<^sub>D (d_trg \<star>\<^sub>D F g) \<star>\<^sub>D F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src"
proof -
have "... \<cong>\<^sub>D d_trg \<star>\<^sub>D F g \<star>\<^sub>D F f\<^sup>*\<^sup>C \<star>\<^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 "... \<cong>\<^sub>D (d_trg \<star>\<^sub>D F g) \<star>\<^sub>D F f\<^sup>*\<^sup>C \<star>\<^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 "... \<cong>\<^sub>D g' \<star>\<^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' \<star>\<^sub>D f'\<^sup>*\<^sup>D)"
by simp
thus "\<exists>f' g'. D.is_left_adjoint f' \<and> D.is_left_adjoint g' \<and> r' \<cong>\<^sub>D g' \<star>\<^sub>D f'\<^sup>*\<^sup>D"
using f' g' by auto
qed
show "\<And>f f' \<mu> \<mu>'. \<lbrakk> D.is_left_adjoint f; D.is_left_adjoint f';
\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>D f'\<guillemotright>; \<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>D f'\<guillemotright> \<rbrakk> \<Longrightarrow> D.iso \<mu> \<and> D.iso \<mu>' \<and> \<mu> = \<mu>'"
proof -
fix f f' \<mu> \<mu>'
assume f: "D.is_left_adjoint f"
and f': "D.is_left_adjoint f'"
and \<mu>: "\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>D f'\<guillemotright>"
and \<mu>': "\<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>D f'\<guillemotright>"
have "C.is_left_adjoint (F.right_map f) \<and> C.is_left_adjoint (F.right_map f')"
using f f' E.G.preserves_left_adjoint by blast
moreover have "\<guillemotleft>F.right_map \<mu> : F.right_map f \<Rightarrow>\<^sub>C F.right_map f'\<guillemotright> \<and>
\<guillemotleft>F.right_map \<mu>' : F.right_map f \<Rightarrow>\<^sub>C F.right_map f'\<guillemotright>"
using \<mu> \<mu>' E.G.preserves_hom by simp
ultimately have "C.iso (F.right_map \<mu>) \<and> C.iso (F.right_map \<mu>') \<and>
F.right_map \<mu> = F.right_map \<mu>'"
using C.BS3 by blast
thus "D.iso \<mu> \<and> D.iso \<mu>' \<and> \<mu> = \<mu>'"
using \<mu> \<mu>' G.reflects_iso G.is_faithful by blast
qed
show "\<And>f g. \<lbrakk> D.is_left_adjoint f; D.is_left_adjoint g; src\<^sub>D f = src\<^sub>D g \<rbrakk>
\<Longrightarrow> \<exists>r \<rho>. tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D r \<rho> 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: "\<exists>r \<rho>. tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> (F.right_map f) (F.right_map g)"
using C.BS2 by simp
obtain r \<rho> where
\<rho>: "tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> (F.right_map f) (F.right_map g)"
using 1 by auto
interpret \<rho>: tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> \<open>F.right_map f\<close> \<open>F.right_map g\<close>
using \<rho> by simp
obtain r' where
r': "D.ide r' \<and> D.in_hhom r' (trg\<^sub>D f) (trg\<^sub>D g) \<and> C.isomorphic (F.right_map r') r"
using f g \<rho>.ide_base \<rho>.tab_in_hom G.locally_essentially_surjective
by (metis D.obj_trg E.G.preserves_reflects_arr E.G.preserves_trg \<rho>.leg0_simps(2-3)
\<rho>.leg1_simps(2,4) \<rho>.base_in_hom(1))
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : r \<Rightarrow>\<^sub>C F.right_map r'\<guillemotright> \<and> C.iso \<phi>"
using r' C.isomorphic_symmetric by blast
have \<sigma>: "tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
(F.right_map r') ((\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho>) (F.right_map f) (F.right_map g)"
using \<phi> \<rho>.is_preserved_by_base_iso by simp
have 1: "\<exists>\<rho>'. \<guillemotleft>\<rho>' : g \<Rightarrow>\<^sub>D H\<^sub>D r' f\<guillemotright> \<and>
F.right_map \<rho>' = F.right_cmp (r', f) \<cdot>\<^sub>C (\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho>"
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 "\<guillemotleft>F.right_cmp (r', f) \<cdot>\<^sub>C (\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho> :
F.right_map g \<Rightarrow>\<^sub>C F.right_map (r' \<star>\<^sub>D f)\<guillemotright>"
using f g r' \<phi> D.left_adjoint_is_ide \<rho>.ide_base
by (intro C.comp_in_homI, auto)
ultimately show ?thesis
using G.locally_full by simp
qed
obtain \<rho>' where \<rho>': "\<guillemotleft>\<rho>' : g \<Rightarrow>\<^sub>D H\<^sub>D r' f\<guillemotright> \<and>
F.right_map \<rho>' = F.right_cmp (r', f) \<cdot>\<^sub>C (\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho>"
using 1 by auto
have "tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D r' \<rho>' f g"
proof -
have "C.inv (F.right_cmp (r', f)) \<cdot>\<^sub>C F.right_map \<rho>' = (\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho>"
using r' f \<rho>' 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 \<sigma> \<rho>' G.reflects_tabulation
by (simp add: D.left_adjoint_is_ide f r')
qed
thus "\<exists>r' \<rho>'. tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D r' \<rho>' f g"
by auto
qed
qed
qed
subsection "Span(C) is a Bicategory of Spans"
text \<open>
We first consider an arbitrary 1-cell \<open>r\<close> in \<open>Span(C)\<close>, 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).
\<close>
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 "\<cdot>" 55)
and prj0 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<p>\<^sub>0[_, _]")
and prj1 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<p>\<^sub>1[_, _]")
and r :: "'a arrow_of_spans_data"
begin
text \<open>
CKS say: ``Suppose \<open>r = (r\<^sub>0, R, r\<^sub>1): A \<rightarrow> B\<close> and put \<open>f = (1, R, r\<^sub>0)\<close>, \<open>g = (1, R, r\<^sub>1)\<close>.
Let \<open>k\<^sub>0, k\<^sub>1\<close> form a kernel pair for \<open>r\<^sub>0\<close> and define \<open>\<rho>\<close> by \<open>k\<^sub>0\<rho> = k\<^sub>1\<rho> = 1\<^sub>R\<close>.''
\<close>
abbreviation ra where "ra \<equiv> r.dom.apex"
abbreviation r0 where "r0 \<equiv> r.dom.leg0"
abbreviation r1 where "r1 \<equiv> r.dom.leg1"
abbreviation f where "f \<equiv> mkIde ra r0"
abbreviation g where "g \<equiv> mkIde ra r1"
abbreviation k0 where "k0 \<equiv> \<p>\<^sub>0[r0, r0]"
abbreviation k1 where "k1 \<equiv> \<p>\<^sub>1[r0, r0]"
text \<open>
Here \<open>ra\<close> is the apex \<open>R\<close> of the span \<open>(r\<^sub>0, R, r\<^sub>1)\<close>, and the spans \<open>f\<close> and \<open>g\<close> also have
that same 0-cell as their apex. The tabulation 2-cell \<open>\<rho>\<close> has to be an arrow of spans
from \<open>g = (1, R, r\<^sub>1)\<close> to \<open>r \<star> f\<close>, which is the span \<open>(k\<^sub>0, r\<^sub>1 \<cdot> k\<^sub>1)\<close>.
\<close>
abbreviation \<rho> where "\<rho> \<equiv> \<lparr>Chn = \<langle>ra \<lbrakk>r0, r0\<rbrakk> ra\<rangle>,
Dom = \<lparr>Leg0 = ra, Leg1 = r1\<rparr>,
Cod = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<rparr>"
lemma has_tabulation:
shows "tabulation vcomp hcomp assoc unit src trg r \<rho> 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 \<open>r \<star> f\<close>
using rf.ide_composite ide_char' by auto
let ?rf = "r \<star> f"
(* TODO: Put this expansion into two_composable_identity_arrows_of_spans. *)
have rf: "?rf = \<lparr>Chn = r0 \<down>\<down> r0,
Dom = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>,
Cod = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<rparr>"
unfolding hcomp_def chine_hcomp_def
using hseq_rf C.comp_cod_arr by auto
interpret Cod_rf: span_in_category C \<open>\<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<close>
using ide_r ide_f rf C.comp_cod_arr
by unfold_locales auto
have Dom_g: "Dom g = \<lparr>Leg0 = ra, Leg1 = r1\<rparr>" by simp
interpret Dom_g: span_in_category C \<open>\<lparr>Leg0 = ra, Leg1 = r1\<rparr>\<close>
using Dom_g g.dom.span_in_category_axioms by simp
interpret Dom_\<rho>: span_in_category C \<open>Dom \<rho>\<close>
using Dom_g g.dom.span_in_category_axioms by simp
interpret Cod_\<rho>: span_in_category C \<open>Cod \<rho>\<close>
using rf Cod_rf.span_in_category_axioms by simp
interpret \<rho>: arrow_of_spans C \<rho>
using Dom_\<rho>.apex_def Cod_\<rho>.apex_def C.comp_assoc C.comp_arr_dom
by unfold_locales auto
have \<rho>: "\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>"
using rf ide_g arr_char dom_char cod_char \<rho>.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 \<rho> f g"
proof -
interpret T: tabulation_data vcomp hcomp assoc unit src trg r \<rho> f g
using ide_f \<rho> by unfold_locales auto
show ?thesis
proof
show T1: "\<And>u \<omega>.
\<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof -
fix u \<omega>
assume u: "ide u"
assume \<omega>: "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
show "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof -
interpret u: identity_arrow_of_spans C u
using u ide_char' by auto
have v: "ide (dom \<omega>)"
using \<omega> by auto
interpret v: identity_arrow_of_spans C \<open>dom \<omega>\<close>
using v ide_char' by auto
interpret \<omega>: arrow_of_spans C \<omega>
using \<omega> arr_char by auto
have hseq_ru: "hseq r u"
using u \<omega> 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 \<open>
CKS say:
``We must show that \<open>(f, \<rho>, g)\<close> is a wide tabulation of \<open>r\<close>.
Take \<open>u = (u\<^sub>0, U, u\<^sub>1): X \<rightarrow> A\<close>, \<open>v = (v\<^sub>0, V, v\<^sub>1): X \<rightarrow> B\<close>,
\<open>\<omega>: v \<Rightarrow> ru\<close> as in \<open>T1\<close>. Let \<open>P\<close> be the pullback of \<open>u\<^sub>1, r\<^sub>0\<close>.
Let \<open>w = (v\<^sub>0, V, p\<^sub>1\<omega>): X \<rightarrow> R\<close>, \<open>\<theta> = p\<^sub>0\<omega>: fw \<Rightarrow> u\<close>,
\<open>\<nu> = 1: v \<Rightarrow> gw\<close>; so \<open>\<omega> = (r\<theta>)(\<rho>w)\<nu>\<close> as required.''
\<close>
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 ?\<omega> = "\<omega>.chine"
let ?P = "r0 \<down>\<down> ?u1"
let ?p0 = "\<p>\<^sub>0[r0, ?u1]"
let ?p1 = "\<p>\<^sub>1[r0, ?u1]"
let ?w1 = "?p1 \<cdot> ?\<omega>"
define w where "w = mkIde ?v0 ?w1"
let ?Q = "?R \<down>\<down> ?w1"
let ?q1 = "\<p>\<^sub>1[?R, ?w1]"
let ?\<rho> = "\<langle>?R \<lbrakk>r0, r0\<rbrakk> ?R\<rangle>"
have P: "?P = ru.apex"
using ru.apex_composite by auto
have Chn_\<omega>: "\<guillemotleft>?\<omega> : ?V \<rightarrow>\<^sub>C ?P\<guillemotright>"
using P Chn_in_hom \<omega> by force
have p0\<omega>: "\<guillemotleft>?p0 \<cdot> ?\<omega> : ?V \<rightarrow>\<^sub>C ?U\<guillemotright>"
using Chn_\<omega> ru.legs_form_cospan by auto
have w1: "\<guillemotleft>?w1 : ?V \<rightarrow>\<^sub>C ?R\<guillemotright>"
using Chn_\<omega> ru.legs_form_cospan r.dom.apex_def by blast
have r1w1: "r1 \<cdot> ?w1 = ?v1"
by (metis C.comp_assoc T.base_simps(3) \<omega> \<omega>.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 \<omega> 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: "\<guillemotleft>\<a>[r, f, w] : (r \<star> f) \<star> w \<Rightarrow> r \<star> f \<star> w\<guillemotright>"
using fw.composable ide_f ide_r w rf.composable by auto
have fw_apex_eq: "fw.apex = ?R \<down>\<down> ?w1"
using w_def fw.dom.apex_def hcomp_def hseq_fw hseq_char \<omega> 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 \<down>\<down> ?w1"
using w_def \<omega> w1 gw.dom.apex_def hcomp_def hseq_gw hseq_char by auto
text \<open>
Well, CKS say take \<open>\<theta> = p\<^sub>0\<omega>\<close>, but taking this literally and trying to define
\<open>\<theta>\<close> so that \<open>Chn \<theta> = ?p\<^sub>0 \<cdot> ?\<omega>\<close>, does not yield the required \<open>dom \<theta> = ?R \<down>\<down> ?w1\<close>.
We need \<open>\<guillemotleft>Chn \<theta> : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C ?U\<guillemotright>\<close>, so we have to compose with a
projection, which leads to: \<open>Chn \<theta> = ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]\<close>.
What CKS say is only correct if the projection \<open>\<p>\<^sub>0[?R, ?w1]\<close> is an identity,
which can't be guaranteed for an arbitrary choice of pullbacks.
\<close>
define \<theta>
where
"\<theta> = \<lparr>Chn = ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1], Dom = Dom (f \<star> w), Cod = Cod u\<rparr>"
interpret Dom_\<theta>: span_in_category C \<open>Dom \<theta>\<close>
using \<theta>_def fw.dom.span_in_category_axioms by simp
interpret Cod_\<theta>: span_in_category C \<open>Cod \<theta>\<close>
using \<theta>_def u.cod.span_in_category_axioms by simp
have Dom_\<theta>_leg0_eq: "Dom_\<theta>.leg0 = ?v0 \<cdot> \<p>\<^sub>0[?R, ?w1]"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Dom_\<theta>_leg1_eq: "Dom_\<theta>.leg1 = r0 \<cdot> ?q1"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Cod_\<theta>_leg0_eq: "Cod_\<theta>.leg0 = ?u0"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Cod_\<theta>_leg1_eq: "Cod_\<theta>.leg1 = ?u1"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Chn_\<theta>_eq: "Chn \<theta> = ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]"
using \<theta>_def by simp
interpret \<theta>: arrow_of_spans C \<theta>
proof
show 1: "\<guillemotleft>Chn \<theta> : Dom_\<theta>.apex \<rightarrow>\<^sub>C Cod_\<theta>.apex\<guillemotright>"
using \<theta>_def Chn_\<omega> ru.legs_form_cospan fw_apex_eq
by (intro C.in_homI, auto)
show "Cod_\<theta>.leg0 \<cdot> Chn \<theta> = Dom_\<theta>.leg0"
proof -
have "Cod_\<theta>.leg0 \<cdot> Chn \<theta> = ?u0 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]"
using Cod_\<theta>_leg0_eq Chn_\<theta>_eq by simp
also have "... = ?v0 \<cdot> \<p>\<^sub>0[?R, ?w1]"
proof -
have "?u0 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1] = (?u0 \<cdot> ?p0 \<cdot> ?\<omega>) \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.comp_assoc by simp
also have "... = ?v0 \<cdot> \<p>\<^sub>0[?R, ?w1]"
proof -
have "?u0 \<cdot> ?p0 \<cdot> ?\<omega> = (?u0 \<cdot> ?p0) \<cdot> ?\<omega>"
using C.comp_assoc by simp
also have "... = \<omega>.cod.leg0 \<cdot> ?\<omega>"
by (metis \<omega> arrow_of_spans_data.select_convs(2) cod_char in_homE
ru.leg0_composite)
also have "... = \<omega>.dom.leg0"
using \<omega>.leg0_commutes by simp
also have "... = ?v0"
using \<omega> dom_char by auto
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
also have "... = Dom_\<theta>.leg0"
using Dom_\<theta>_leg0_eq by simp
finally show ?thesis by blast
qed
show "Cod_\<theta>.leg1 \<cdot> Chn \<theta> = Dom_\<theta>.leg1"
proof -
have "Cod_\<theta>.leg1 \<cdot> Chn \<theta> = ?u1 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]"
using Cod_\<theta>_leg1_eq Chn_\<theta>_eq by simp
also have "... = r0 \<cdot> ?q1"
proof -
have "?u1 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1] = ((?u1 \<cdot> ?p0) \<cdot> ?\<omega>) \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.comp_assoc by fastforce
also have "... = ((r0 \<cdot> ?p1) \<cdot> ?\<omega>) \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.pullback_commutes' ru.legs_form_cospan by simp
also have "... = r0 \<cdot> ?w1 \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.comp_assoc by fastforce
also have "... = r0 \<cdot> ?R \<cdot> ?q1"
using \<omega> C.in_homE C.pullback_commutes' w1 by auto
also have "... = r0 \<cdot> ?q1"
using \<omega> w1 C.comp_cod_arr by auto
finally show ?thesis by simp
qed
also have "... = Dom_\<theta>.leg1"
using Dom_\<theta>_leg1_eq by simp
finally show ?thesis by blast
qed
qed
text \<open>
Similarly, CKS say to take \<open>\<nu> = 1: v \<Rightarrow> gw\<close>, but obviously this can't be
interpreted literally, either, because \<open>v.apex\<close> and \<open>gw.apex\<close> are not equal.
Instead, we have to define \<open>\<nu>\<close> so that \<open>Chn \<nu> = C.inv \<p>\<^sub>0[?R, ?w1]\<close>,
noting that \<open>\<p>\<^sub>0[?R, ?w1]\<close> is the pullback of an identity,
hence is an isomorphism. Then \<open>?v0 \<cdot> \<p>\<^sub>0[?R, ?w1] \<cdot> Chn \<nu> = ?v0\<close> and
\<open>?v1 \<cdot> \<p>\<^sub>1[?R, ?w1] \<cdot> Chn \<nu> = ?v1 \<cdot> ?w1\<close>, showing that \<open>\<nu>\<close> is an arrow
of spans.
\<close>
let ?\<nu>' = "\<p>\<^sub>0[?R, ?w1]"
define \<nu>
where
"\<nu> = \<lparr>Chn = C.inv ?\<nu>', Dom = Dom (dom \<omega>), Cod = Cod (g \<star> w)\<rparr>"
have iso_\<nu>: "C.inverse_arrows ?\<nu>' (Chn \<nu>)"
using \<nu>_def \<omega> 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 \<open>
$$
\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} \\
}
$$
\<close>
have w1_eq: "?w1 = ?q1 \<cdot> C.inv ?\<nu>'"
proof -
have "?R \<cdot> ?q1 = ?w1 \<cdot> ?\<nu>'"
using iso_\<nu> \<omega> w1 C.pullback_commutes [of ?R ?w1] by blast
hence "?q1 = ?w1 \<cdot> ?\<nu>'"
using \<omega> w1 C.comp_cod_arr by auto
thus ?thesis
using iso_\<nu> 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_\<nu>: span_in_category C \<open>Dom \<nu>\<close>
using \<nu>_def v.dom.span_in_category_axioms by simp
interpret Cod_\<nu>: span_in_category C \<open>Cod \<nu>\<close>
using \<nu>_def gw.cod.span_in_category_axioms by simp
interpret \<nu>: arrow_of_spans C \<nu>
proof
show 1: "\<guillemotleft>Chn \<nu> : Dom_\<nu>.apex \<rightarrow>\<^sub>C Cod_\<nu>.apex\<guillemotright>"
proof
show "C.arr (Chn \<nu>)"
using iso_\<nu> by auto
show "C.dom (Chn \<nu>) = Dom_\<nu>.apex"
using \<nu>_def iso_\<nu> w1 gw_apex_eq by fastforce
show "C.cod (Chn \<nu>) = Cod_\<nu>.apex"
using \<nu>_def iso_\<nu> 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_\<nu>.leg0 \<cdot> Chn \<nu> = Dom_\<nu>.leg0"
using w_def \<nu>_def 1 iso_\<nu> hcomp_def hseq_gw C.comp_arr_inv
C.comp_assoc v.leg0_commutes
by auto
show "Cod_\<nu>.leg1 \<cdot> Chn \<nu> = Dom_\<nu>.leg1"
using w_def \<nu>_def hcomp_def hseq_gw C.comp_assoc w1_eq r1w1
by auto
qed
text \<open>
Now we can proceed to establishing the conclusions of \<open>T1\<close>.
\<close>
have "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> dom \<rho> \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof (intro conjI)
show ide_w: "ide w"
using w by blast
show \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
using \<theta>_def \<theta>.arrow_of_spans_axioms arr_char dom_char cod_char hseq_fw hseq_char
hcomp_def fw.chine_eq_apex
by auto
show \<nu>: "\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> dom \<rho> \<star> w\<guillemotright>"
proof -
have "\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright>"
using \<nu>_def \<omega> \<nu>.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 \<nu>"
using iso_\<nu> iso_char arr_char \<nu>.arrow_of_spans_axioms by auto
show "T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof (intro arr_eqI)
have \<rho>w: "\<guillemotleft>\<rho> \<star> w : g \<star> w \<Rightarrow> (r \<star> f) \<star> w\<guillemotright>"
using w_def \<rho> ide_w hseq_rf hseq_fw hseq_gw by auto
have r\<theta>: "\<guillemotleft>r \<star> \<theta> : r \<star> f \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using arfw ide_r \<theta> fw.composite_simps(2) rf.composable by auto
have 1: "\<guillemotleft>T.composite_cell w \<theta> \<bullet> \<nu> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
using \<nu> \<rho>w arfw r\<theta> by auto
show 3: "par (T.composite_cell w \<theta> \<bullet> \<nu>) \<omega>"
using 1 \<omega> by (elim in_homE, auto)
show "Chn (T.composite_cell w \<theta> \<bullet> \<nu>) = ?\<omega>"
proof -
have 2: "Chn (T.composite_cell w \<theta> \<bullet> \<nu>) =
Chn (r \<star> \<theta>) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using 1 3 Chn_vcomp C.comp_assoc
by (metis (full_types) seqE)
also have "... = ?\<omega>"
proof -
let ?LHS = "Chn (r \<star> \<theta>) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
have Chn_r\<theta>: "Chn (r \<star> \<theta>) = \<langle>r.chine \<cdot> \<p>\<^sub>1[r0, r0 \<cdot> ?q1]
\<lbrakk>r0, ?u1\<rbrakk>
\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?q1]\<rangle>"
using r\<theta> hcomp_def \<theta>_def chine_hcomp_def Dom_\<theta>_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 \<a>[r, f, w] = rfw.chine_assoc"
using \<alpha>_ide ide_f rf.composable fw.composable w by auto
have Chn_\<rho>w: "Chn (\<rho> \<star> w) = \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
proof -
have "Chn (\<rho> \<star> w) =
chine_hcomp
\<lparr>Chn = ?\<rho>,
Dom = \<lparr>Leg0 = ?R, Leg1 = r1\<rparr>,
Cod = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<rparr>
\<lparr>Chn = v.apex,
Dom = \<lparr>Leg0 = ?v0, Leg1 = ?w1\<rparr>,
Cod = \<lparr>Leg0 = ?v0, Leg1 = ?w1\<rparr>\<rparr>"
using \<rho> ide_w hseq_rf hseq_char hcomp_def src_def trg_def
by (metis (no_types, lifting) \<rho>w arrI arrow_of_spans_data.select_convs(1)
v.dom.apex_def w_def)
also have "... = \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?V \<cdot> ?\<nu>'\<rangle>"
unfolding chine_hcomp_def by simp
also have "... = \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
proof -
have "?V \<cdot> ?\<nu>' = ?\<nu>'"
using C.comp_ide_arr v.dom.ide_apex \<rho> w1 by auto
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
have 3: "C.seq ?p1 ?\<omega>"
using w1 by blast
moreover have 4: "C.seq ?p1 ?LHS"
proof
show "\<guillemotleft>?LHS : v.apex \<rightarrow>\<^sub>C ru.apex\<guillemotright>"
by (metis (no_types, lifting) 1 2 Chn_in_hom ru.chine_eq_apex
v.chine_eq_apex)
show "\<guillemotleft>?p1 : ru.apex \<rightarrow>\<^sub>C ?R\<guillemotright>"
using P C.prj1_in_hom ru.legs_form_cospan by fastforce
qed
moreover have "?p0 \<cdot> ?LHS = ?p0 \<cdot> ?\<omega>"
proof -
have "?p0 \<cdot> ?LHS =
(?p0 \<cdot> Chn (r \<star> \<theta>)) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using C.comp_assoc by simp
also have "... = (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?q1]) \<cdot>
Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
proof -
have "?p0 \<cdot> Chn (r \<star> \<theta>) = \<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?q1]"
by (metis C.prj_tuple(1) Chn_r\<theta> \<theta>_def arrI Dom_\<theta>_leg1_eq
arrow_of_spans_data.select_convs(3) chine_hcomp_props(2)
hseq_char r.cod_simps(2) r\<theta> u.cod_simps(3))
thus ?thesis by argo
qed
also have
"... = ?p0 \<cdot> ?\<omega> \<cdot> (rfw.Prj\<^sub>0\<^sub>0 \<cdot> Chn \<a>[r, f, w]) \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using w_def \<theta>_def C.comp_assoc by simp
also have "... = ?p0 \<cdot> ?\<omega> \<cdot> (rfw.Prj\<^sub>0 \<cdot> Chn (\<rho> \<star> w)) \<cdot> Chn \<nu>"
using Chn_arfw rfw.prj_chine_assoc C.comp_assoc by simp
also have "... = ?p0 \<cdot> ?\<omega> \<cdot> ?\<nu>' \<cdot> Chn \<nu>"
proof -
have "rfw.Prj\<^sub>0 \<cdot> Chn (\<rho> \<star> w) = \<p>\<^sub>0[k0, ?w1] \<cdot> \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
using w_def Chn_\<rho>w C.comp_cod_arr by simp
also have "... = ?\<nu>'"
by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(1) C.seqE
C.tuple_is_extensional Chn_\<rho>w 4)
finally have "rfw.Prj\<^sub>0 \<cdot> Chn (\<rho> \<star> w) = ?\<nu>'"
by blast
thus ?thesis by simp
qed
also have "... = ?p0 \<cdot> ?\<omega>"
using iso_\<nu> C.comp_arr_dom
by (metis (no_types, lifting) C.comp_arr_inv C.dom_comp \<nu>_def
\<omega>.chine_simps(1) 3 arrow_of_spans_data.simps(1) w1_eq)
finally show ?thesis by blast
qed
moreover have "?p1 \<cdot> ?LHS = ?w1"
proof -
have "?p1 \<cdot> ?LHS =
(?p1 \<cdot> Chn (r \<star> \<theta>)) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using C.comp_assoc by simp
also have "... = (r.chine \<cdot> \<p>\<^sub>1[r0, r0 \<cdot> ?q1]) \<cdot> Chn \<a>[r, f, w] \<cdot>
Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2) C.seqE
C.tuple_is_extensional Chn_r\<theta> 4)
also have "... = r.chine \<cdot> (rfw.Prj\<^sub>1 \<cdot> Chn \<a>[r, f, w]) \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using w_def Dom_\<theta>_leg1_eq C.comp_assoc by simp
also have "... = r.chine \<cdot> (rfw.Prj\<^sub>1\<^sub>1 \<cdot> Chn (\<rho> \<star> w)) \<cdot> Chn \<nu>"
using Chn_arfw rfw.prj_chine_assoc(1) C.comp_assoc by simp
also have "... = r.chine \<cdot> ?q1 \<cdot> Chn \<nu>"
proof -
have "rfw.Prj\<^sub>1\<^sub>1 \<cdot> Chn (\<rho> \<star> w) =
(k1 \<cdot> \<p>\<^sub>1[k0, ?w1]) \<cdot> \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
using w_def Chn_\<rho>w C.comp_cod_arr by simp
also have "... = k1 \<cdot> \<p>\<^sub>1[k0, ?w1] \<cdot> \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
using C.comp_assoc by simp
also have "... = k1 \<cdot> ?\<rho> \<cdot> ?q1"
by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2)
C.seqE C.tuple_is_extensional Chn_\<rho>w 4)
also have "... = (k1 \<cdot> ?\<rho>) \<cdot> ?q1"
using C.comp_assoc by presburger
also have "... = ?R \<cdot> ?q1"
by simp
also have "... = ?q1"
by (metis Dom_\<theta>_leg1_eq C.comp_ide_arr C.prj1_simps(3)
C.prj1_simps_arr C.seqE C.seqI Dom_\<theta>.leg_simps(3)
r.dom.ide_apex)
finally have "rfw.Prj\<^sub>1\<^sub>1 \<cdot> Chn (\<rho> \<star> w) = ?q1"
by blast
thus ?thesis by simp
qed
also have "... = (r.chine \<cdot> ?p1) \<cdot> ?\<omega>"
using \<nu>_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: "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w';
\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>;
T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<bullet> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume ide_w: "ide w"
assume ide_w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
assume E: "T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<bullet> \<beta>"
interpret T: uw\<theta>w'\<theta>'\<beta> vcomp hcomp assoc unit src trg r \<rho> f g u w \<theta> w' \<theta>' \<beta>
using ide_w ide_w' \<theta> \<theta>' \<beta> E comp_assoc
by unfold_locales auto
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>)"
proof
interpret u: identity_arrow_of_spans C u
using T.uw\<theta>.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 = "\<p>\<^sub>0[?R, ?w1]"
let ?p0' = "\<p>\<^sub>0[?R, ?w1']"
let ?p1 = "\<p>\<^sub>1[?R, ?w1]"
let ?p1' = "\<p>\<^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 \<open>r \<star> f \<star> w\<close>
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 \<open>r \<star> f \<star> w'\<close>
using rfw'.composites_are_identities ide_char' by auto
have \<rho>w: "\<guillemotleft>\<rho> \<star> w : g \<star> w \<Rightarrow> (r \<star> f) \<star> w\<guillemotright>"
using \<rho> hseq_gw by blast
interpret \<rho>w: two_composable_arrows_of_spans C prj0 prj1 \<rho> w
using \<rho>w by unfold_locales auto
have \<rho>w': "\<guillemotleft>\<rho> \<star> w' : g \<star> w' \<Rightarrow> (r \<star> f) \<star> w'\<guillemotright>"
using \<rho> hseq_gw' by blast
interpret \<rho>w': two_composable_arrows_of_spans C prj0 prj1 \<rho> w'
using \<rho>w' by unfold_locales auto
have arfw: "\<guillemotleft>\<a>[r, f, w] : (r \<star> f) \<star> w \<Rightarrow> r \<star> f \<star> w\<guillemotright>"
using fw.composable ide_f ide_r ide_w rf.composable by auto
have arfw': "\<guillemotleft>\<a>[r, f, w'] : (r \<star> f) \<star> w' \<Rightarrow> r \<star> f \<star> w'\<guillemotright>"
using fw'.composable ide_f ide_r ide_w' rf.composable by auto
have r\<theta>: "\<guillemotleft>r \<star> \<theta> : r \<star> f \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
by fastforce
interpret Dom_\<theta>: span_in_category C \<open>Dom \<theta>\<close>
using fw.dom.span_in_category_axioms
by (metis \<theta> arrow_of_spans_data.select_convs(2) in_homE dom_char)
interpret Cod_\<theta>: span_in_category C \<open>Cod \<theta>\<close>
using \<theta> u.cod.span_in_category_axioms cod_char by auto
interpret \<theta>: arrow_of_spans C \<theta>
using arr_char T.uw\<theta>.\<theta>_simps(1) by auto
interpret r\<theta>: two_composable_arrows_of_spans C prj0 prj1 r \<theta>
using r\<theta> by unfold_locales auto
have r\<theta>': "\<guillemotleft>r \<star> \<theta>' : r \<star> f \<star> w' \<Rightarrow> r \<star> u\<guillemotright>"
by fastforce
interpret Dom_\<theta>': span_in_category C \<open>Dom \<theta>'\<close>
using fw'.dom.span_in_category_axioms
by (metis \<theta>' arrow_of_spans_data.select_convs(2) in_homE dom_char)
interpret Cod_\<theta>': span_in_category C \<open>Cod \<theta>'\<close>
using \<theta>' u.cod.span_in_category_axioms cod_char by auto
interpret \<theta>': arrow_of_spans C \<theta>'
using arr_char T.uw'\<theta>'.\<theta>_simps(1) by auto
interpret r\<theta>': two_composable_arrows_of_spans C prj0 prj1 r \<theta>'
using r\<theta>' by unfold_locales auto
have 7: "\<guillemotleft>T.composite_cell w' \<theta>' \<bullet> \<beta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using \<beta> \<rho>w' arfw' r\<theta>' by auto
have 8: "\<guillemotleft>T.composite_cell w \<theta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using \<rho>w arfw r\<theta> by auto
interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u
using hseq_char by unfold_locales auto
interpret Dom_\<beta>: span_in_category C \<open>Dom \<beta>\<close>
using \<beta> 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_\<beta>: span_in_category C \<open>Cod \<beta>\<close>
using \<beta> 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 \<beta>: arrow_of_spans C \<beta>
using \<beta> arr_char by auto
text \<open>
CKS say: ``Take \<open>u\<close>, \<open>w\<close>, \<open>w'\<close>, \<open>\<theta>\<close>, \<open>\<theta>'\<close> as in \<open>T2\<close> and note that
\<open>fw = (w\<^sub>0, W, r\<^sub>0 w\<^sub>1)\<close>, \<open>gw = (w\<^sub>0, W, r\<^sub>1 w\<^sub>1)\<close>, \emph{etc}.
So \<open>\<beta>: W \<rightarrow> W'\<close> satisfies \<open>w\<^sub>0 = w\<^sub>0' \<beta>\<close>, \<open>r\<^sub>1 w\<^sub>1 = r\<^sub>1 w\<^sub>1' \<beta>\<close>.
But the equation \<open>(r\<theta>)(\<rho>w) = (r\<theta>')(\<rho>w')\<beta>\<close> gives \<open>w\<^sub>1 = w\<^sub>1'\<close>.
So \<open>\<gamma> = \<beta> : w \<Rightarrow> w'\<close> is unique with \<open>\<beta> = g \<gamma>, \<theta> = \<theta>' (f \<gamma>).\<close>''
Once again, there is substantial punning in the proof sketch given by CKS.
We can express \<open>fw\<close> and \<open>gw\<close> almost in the form they indicate, but projections
are required.
\<close>
have cospan: "C.cospan ?R ?w1"
using hseq_char [of \<rho> w] src_def trg_def by auto
have cospan': "C.cospan ?R ?w1'"
using hseq_char [of \<rho> w'] src_def trg_def by auto
have fw: "f \<star> w = \<lparr>Chn = ?R \<down>\<down> ?w1,
Dom = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r0 \<cdot> ?p1\<rparr>,
Cod = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r0 \<cdot> ?p1\<rparr>\<rparr>"
using ide_f hseq_char hcomp_def chine_hcomp_def fw.dom.apex_def cospan
fw.chine_eq_apex
by auto
have gw: "g \<star> w = \<lparr>Chn = ?R \<down>\<down> ?w1,
Dom = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>,
Cod = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>\<rparr>"
using hseq_gw hseq_char hcomp_def chine_hcomp_def gw.dom.apex_def cospan
gw.chine_eq_apex
by auto
have fw': "f \<star> w' = \<lparr>Chn = ?R \<down>\<down> ?w1',
Dom = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r0 \<cdot> ?p1'\<rparr>,
Cod = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r0 \<cdot> ?p1'\<rparr>\<rparr>"
using ide_f hseq_char hcomp_def chine_hcomp_def fw'.dom.apex_def cospan'
fw'.chine_eq_apex
by auto
have gw': "g \<star> w' = \<lparr>Chn = ?R \<down>\<down> ?w1',
Dom = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>,
Cod = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>\<rparr>"
using hseq_gw' hseq_char hcomp_def chine_hcomp_def gw'.dom.apex_def cospan'
gw'.chine_eq_apex
by auto
text \<open>
Note that \<open>?p0\<close> and \<open>?p0'\<close> are only isomorphisms, not identities,
and we have \<open>?p1\<close> (which equals \<open>?w1 \<cdot> ?p0\<close>) and \<open>?p1'\<close> (which equals \<open>?w1' \<cdot> ?p0'\<close>)
in place of \<open>?w1\<close> and \<open>?w1'\<close>.
\<close>
text \<open>
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 \<open>W\<close> and \<open>W'\<close> to denote the apexes of
\<open>w\<close> and \<open>w'\<close>, respectively.
We already have the expressions \<open>?R \<down>\<down> ?w1\<close> and \<open>?R \<down>\<down> ?w1'\<close> for the
apexes of \<open>fw\<close> and \<open>fw'\<close> (which are the same as the apexes of
\<open>gw\<close> and \<open>gw'\<close>, respectively) and we will not use any abbreviation for them.
\<close>
text \<open>
$$
\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}
}
$$
\<close>
have Chn_\<beta>: "\<guillemotleft>\<beta>.chine: ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C ?R \<down>\<down> ?w1'\<guillemotright>"
using gw gw' Chn_in_hom \<beta> gw'.chine_eq_apex gw.chine_eq_apex by force
have \<beta>_eq: "\<beta> = \<lparr>Chn = \<beta>.chine,
Dom = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>,
Cod = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>\<rparr>"
using \<beta> gw gw' dom_char cod_char by auto
have Dom_\<beta>_eq: "Dom \<beta> = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>"
using \<beta> gw gw' dom_char cod_char by auto
have Cod_\<beta>_eq: "Cod \<beta> = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>"
using \<beta> gw gw' dom_char cod_char by auto
have \<beta>0: "?w0 \<cdot> ?p0 = ?w0' \<cdot> ?p0' \<cdot> \<beta>.chine"
using Dom_\<beta>_eq Cod_\<beta>_eq \<beta>.leg0_commutes C.comp_assoc by simp
have \<beta>1: "r1 \<cdot> ?p1 = r1 \<cdot> ?p1' \<cdot> \<beta>.chine"
using Dom_\<beta>_eq Cod_\<beta>_eq \<beta>.leg1_commutes C.comp_assoc by simp
have Dom_\<theta>_0: "Dom_\<theta>.leg0 = ?w0 \<cdot> ?p0"
using arrI dom_char fw T.uw\<theta>.\<theta>_simps(4) by auto
have Cod_\<theta>_0: "Cod_\<theta>.leg0 = ?u0"
using \<theta> cod_char by auto
have Dom_\<theta>_1: "Dom_\<theta>.leg1 = r0 \<cdot> ?p1"
using arrI dom_char fw T.uw\<theta>.\<theta>_simps(4) by auto
have Cod_\<theta>_1: "Cod_\<theta>.leg1 = ?u1"
using T.uw\<theta>.\<theta>_simps(5) cod_char by auto
have Dom_\<theta>'_0: "Dom_\<theta>'.leg0 = ?w0' \<cdot> ?p0'"
using dom_char fw' T.uw'\<theta>'.\<theta>_simps(4) by auto
have Cod_\<theta>'_0: "Cod_\<theta>'.leg0 = ?u0"
using T.uw'\<theta>'.\<theta>_simps(5) cod_char by auto
have Dom_\<theta>'_1: "Dom_\<theta>'.leg1 = r0 \<cdot> ?p1'"
using dom_char fw' T.uw'\<theta>'.\<theta>_simps(4) by auto
have Cod_\<theta>'_1: "Cod_\<theta>'.leg1 = ?u1"
using T.uw'\<theta>'.\<theta>_simps(5) cod_char by auto
have Dom_\<rho>_0: "Dom_\<rho>.leg0 = ?R"
by simp
have Dom_\<rho>_1: "Dom_\<rho>.leg1 = r1"
by simp
have Cod_\<rho>_0: "Cod_\<rho>.leg0 = k0"
by simp
have Cod_\<rho>_1: "Cod_\<rho>.leg1 = r1 \<cdot> k1"
by simp
have Chn_r\<theta>: "\<guillemotleft>r\<theta>.chine : rfw.chine \<rightarrow>\<^sub>C ru.chine\<guillemotright>"
using r\<theta>.chine_composite_in_hom ru.chine_composite rfw.chine_composite
Cod_\<theta>_1 Dom_\<theta>_1 fw.leg1_composite
by auto
have Chn_r\<theta>_eq: "r\<theta>.chine = \<langle>\<p>\<^sub>1[r0, r0 \<cdot> ?p1] \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1]\<rangle>"
using r\<theta>.chine_composite Cod_\<theta>_1 Dom_\<theta>_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\<theta>_cod_apex_eq: "r\<theta>.cod.apex = r0 \<down>\<down> ?u1"
using Cod_\<theta>_1 r\<theta>.chine_composite_in_hom by auto
hence r\<theta>'_cod_apex_eq: "r\<theta>'.cod.apex = r0 \<down>\<down> ?u1"
using Cod_\<theta>'_1 r\<theta>'.chine_composite_in_hom by auto
have Chn_r\<theta>': "\<guillemotleft>r\<theta>'.chine : rfw'.chine \<rightarrow>\<^sub>C ru.chine\<guillemotright>"
using r\<theta>'.chine_composite_in_hom ru.chine_composite rfw'.chine_composite
Cod_\<theta>'_1 Dom_\<theta>'_1 fw'.leg1_composite
by auto
have Chn_r\<theta>'_eq: "r\<theta>'.chine =
\<langle>\<p>\<^sub>1[r0, r0 \<cdot> ?p1'] \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1']\<rangle>"
using r\<theta>'.chine_composite Cod_\<theta>'_1 Dom_\<theta>'_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_\<rho>w: "\<guillemotleft>\<rho>w.chine : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C k0 \<down>\<down> ?w1\<guillemotright>"
using \<rho>w.chine_composite_in_hom by simp
have Chn_\<rho>w_eq: "\<rho>w.chine = \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using \<rho>w.chine_composite C.comp_cod_arr ide_w
by (simp add: chine_hcomp_arr_ide hcomp_def)
have Chn_\<rho>w': "\<guillemotleft>\<rho>w'.chine : ?R \<down>\<down> ?w1' \<rightarrow>\<^sub>C k0 \<down>\<down> ?w1'\<guillemotright>"
using \<rho>w'.chine_composite_in_hom by simp
have Chn_\<rho>w'_eq: "\<rho>w'.chine = \<langle>\<rho>.chine \<cdot> ?p1' \<lbrakk>k0, ?w1'\<rbrakk> ?p0'\<rangle>"
using \<rho>w'.chine_composite C.comp_cod_arr ide_w' Dom_\<rho>_0 Cod_\<rho>_0
by (metis \<rho>w'.composite_is_arrow chine_hcomp_arr_ide chine_hcomp_def hseq_char
w'.cod_simps(3))
text \<open>
The following are some collected commutativity properties that are used
subsequently.
\<close>
have "C.commutative_square r0 ?u1 ?p1 \<theta>.chine"
using ru.legs_form_cospan(1) Dom_\<theta>.is_span Dom_\<theta>_1 Cod_\<theta>_1 \<theta>.leg1_commutes
by (intro C.commutative_squareI) auto
have "C.commutative_square r0 ?u1 (?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)"
by (metis (mono_tags, lifting) C.commutative_square_comp_arr C.dom_comp
C.seqE Cod_\<theta>'_1 Dom_\<beta>.leg_simps(3) Dom_\<beta>_eq Dom_\<theta>'.leg_simps(3)
Dom_\<theta>'_1 \<beta>1 \<theta>'.leg1_commutes C.commutative_squareI
ru.legs_form_cospan(1) span_data.simps(2))
have "C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1] (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1])"
using ru.legs_form_cospan(1) Dom_\<theta>.is_span Dom_\<theta>_1
C.comp_assoc C.pullback_commutes' r\<theta>.legs_form_cospan(1)
apply (intro C.commutative_squareI)
apply auto
by (metis C.comp_assoc Cod_\<theta>_1 \<theta>.leg1_commutes)
hence "C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1] (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1])"
using fw.leg1_composite by auto
have "C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1'] (\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1'])"
using C.tuple_is_extensional Chn_r\<theta>'_eq r\<theta>'.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
by (intro C.commutative_squareI) auto
have "C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0"
by (metis (no_types, lifting) C.commutative_square_comp_arr C.comp_assoc
C.pullback_commutes select_convs(2) rfw'.cospan_\<nu>\<pi>
rfw'.prj_chine_assoc(2) rfw'.prj_chine_assoc(3) rfw'.prj_simps(2)
span_data.select_convs(1))
have "C.commutative_square r0 (r0 \<cdot> ?p1) rfw.Prj\<^sub>1\<^sub>1 \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>"
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 \<cdot> ?p1') rfw'.Prj\<^sub>1\<^sub>1 \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>"
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 (\<rho>.chine \<cdot> ?p1) ?p0"
using C.tuple_is_extensional Chn_\<rho>w_eq \<rho>w.chine_simps(1) by fastforce
have "C.commutative_square k0 ?w1' (\<rho>.chine \<cdot> ?p1') (w'.chine \<cdot> ?p0')"
using C.tuple_is_extensional \<rho>w'.chine_composite \<rho>w'.chine_simps(1) by force
have "C.commutative_square k0 ?w1' (\<rho>.chine \<cdot> ?p1') ?p0'"
using C.tuple_is_extensional Chn_\<rho>w'_eq \<rho>w'.chine_simps(1) by force
text \<open>
Now, derive the consequences of the equation:
\[
\<open>(r \<star> \<theta>) \<bullet> \<a>[r, ?f, w] \<bullet> (?\<rho> \<star> w) = (r \<star> \<theta>') \<bullet> \<a>[r, ?f, w'] \<bullet> (?\<rho> \<star> w') \<bullet> \<beta>\<close>
\]
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.
\<close>
have R: "Chn (T.composite_cell w' \<theta>' \<bullet> \<beta>) =
\<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle>"
proof -
have "Chn (T.composite_cell w' \<theta>' \<bullet> \<beta>) =
r\<theta>'.chine \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
proof -
have 1: "\<guillemotleft>T.composite_cell w' \<theta>' \<bullet> \<beta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using \<beta> \<rho>w' arfw' r\<theta>' by auto
have "Chn (T.composite_cell w' \<theta>' \<bullet> \<beta>) = Chn (T.composite_cell w' \<theta>') \<cdot> \<beta>.chine"
using 1 Chn_vcomp by blast
also have "... = (r\<theta>'.chine \<cdot> Chn (\<a>[r, f, w'] \<bullet> (\<rho> \<star> w'))) \<cdot> \<beta>.chine"
proof -
have "seq (r \<star> \<theta>') (\<a>[r, f, w'] \<bullet> (\<rho> \<star> w'))"
using 1 by blast
thus ?thesis
using 1 Chn_vcomp by presburger
qed
also have "... = (r\<theta>'.chine \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
proof -
have "seq \<a>[r, f, w'] (\<rho> \<star> 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 "... = \<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle>"
proof -
let ?LHS = "r\<theta>'.chine \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
let ?RHS = "\<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle>"
have LHS: "\<guillemotleft>?LHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r\<theta>'.cod.apex\<guillemotright>"
proof (intro C.comp_in_homI)
show "\<guillemotleft>\<beta>.chine : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C ?R \<down>\<down> ?w1'\<guillemotright>"
using Chn_\<beta> by simp
show "\<guillemotleft>\<rho>w'.chine : ?R \<down>\<down> ?w1' \<rightarrow>\<^sub>C Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1\<guillemotright>"
using Chn_\<rho>w' by simp
show "\<guillemotleft>Chn \<a>[r, f, w'] : Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1 \<rightarrow>\<^sub>C rfw'.chine\<guillemotright>"
using arfw'
by (metis (no_types, lifting) Chn_in_hom Cod_\<rho>_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 "\<guillemotleft>r\<theta>'.chine : rfw'.chine \<rightarrow>\<^sub>C r\<theta>'.cod.apex\<guillemotright>"
using Chn_r\<theta>' by auto
qed
have 2: "C.commutative_square r0 ?u1
(?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)"
by fact
have RHS: "\<guillemotleft>?RHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r\<theta>'.cod.apex\<guillemotright>"
using 2 Chn_\<beta> r\<theta>'_cod_apex_eq
C.tuple_in_hom [of r0 ?u1 "?p1' \<cdot> \<beta>.chine" "\<theta>'.chine \<cdot> \<beta>.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\<theta>'_cod_apex_eq by auto
show "C.seq ru.prj\<^sub>1 ?RHS"
using RHS r\<theta>'_cod_apex_eq by auto
show "ru.prj\<^sub>0 \<cdot> ?LHS = ru.prj\<^sub>0 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>0 \<cdot> ?LHS =
(ru.prj\<^sub>0 \<cdot> r\<theta>'.chine) \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = ((\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1']) \<cdot> Chn \<a>[r, f, w']) \<cdot>
\<rho>w'.chine \<cdot> \<beta>.chine"
using Chn_r\<theta>'_eq C.comp_assoc fw'
\<open>C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1']
(\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1'])\<close>
by simp
also have "... = \<theta>'.chine \<cdot> (\<p>\<^sub>0[r0, r0 \<cdot> ?p1'] \<cdot> Chn \<a>[r, f, w']) \<cdot>
\<rho>w'.chine \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = \<theta>'.chine \<cdot> (\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine) \<cdot>
\<beta>.chine"
using ide_f hseq_rf hseq_char \<alpha>_ide C.comp_assoc
rfw'.chine_assoc_def fw'.leg1_composite C.prj_tuple(1)
\<open>C.commutative_square r0 (r0 \<cdot> ?p1')
rfw'.Prj\<^sub>1\<^sub>1 \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>\<close>
by simp
also have "... = \<theta>'.chine \<cdot> \<beta>.chine"
proof -
have "\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine = gw'.apex"
proof (intro C.prj_joint_monic
[of ?R ?w1' "\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine"
gw'.apex])
show "C.cospan ?R ?w1'"
using fw'.legs_form_cospan(1) by simp
show "C.seq ?p1' (\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine)"
proof (intro C.seqI' C.comp_in_homI)
show "\<guillemotleft>\<rho>w'.chine : Dom_\<rho>.leg0 \<down>\<down> w'.leg1 \<rightarrow>\<^sub>C Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1\<guillemotright>"
using \<rho>w'.chine_composite_in_hom by simp
show "\<guillemotleft>\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, w'.leg1\<rbrakk> rfw'.Prj\<^sub>0\<rangle> :
Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1 \<rightarrow>\<^sub>C ?R \<down>\<down> w'.leg1\<guillemotright>"
using \<open>C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\<close>
C.tuple_in_hom [of ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0]
rf rf.leg0_composite
by auto
show "\<guillemotleft>?p1' : ?R \<down>\<down> w'.leg1 \<rightarrow>\<^sub>C f.apex\<guillemotright>"
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' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
?p0' \<cdot> gw'.apex"
proof -
have "?p0' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
(?p0' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w'.chine"
using C.comp_assoc by simp
also have "... = rfw'.Prj\<^sub>0 \<cdot> \<rho>w'.chine"
using \<open>C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\<close> by auto
also have
"... = \<p>\<^sub>0[k0, ?w1'] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1' \<lbrakk>k0, ?w1'\<rbrakk> w'.chine \<cdot> ?p0'\<rangle>"
using \<rho>w'.chine_composite Dom_\<rho>_0 Cod_\<rho>_0 C.comp_cod_arr by simp
also have "... = w'.chine \<cdot> ?p0'"
using \<open>C.commutative_square k0 ?w1'
(\<rho>.chine \<cdot> ?p1') (w'.chine \<cdot> ?p0')\<close>
by simp
also have "... = ?p0' \<cdot> 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' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
?p1' \<cdot> gw'.apex"
proof -
have "?p1' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
(?p1' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w'.chine"
using C.comp_assoc by simp
also have "... = rfw'.Prj\<^sub>0\<^sub>1 \<cdot> \<rho>w'.chine"
using \<open>C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\<close> by simp
also have
"... = k0 \<cdot> \<p>\<^sub>1[k0, ?w1'] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1' \<lbrakk>k0, ?w1'\<rbrakk> w'.chine \<cdot> ?p0'\<rangle>"
using \<rho>w'.chine_composite Cod_\<rho>_0 C.comp_assoc C.comp_cod_arr
by simp
also have "... = k0 \<cdot> \<rho>.chine \<cdot> ?p1'"
using \<open>C.commutative_square k0 ?w1'
(\<rho>.chine \<cdot> ?p1') (w'.chine \<cdot> ?p0')\<close>
by simp
also have "... = (k0 \<cdot> \<rho>.chine) \<cdot> ?p1'"
using C.comp_assoc by metis
also have "... = ?p1'"
using \<rho>.leg0_commutes C.comp_cod_arr cospan' by simp
also have "... = ?p1' \<cdot> 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_\<beta> C.comp_cod_arr gw'.apex_composite by auto
qed
also have "... = \<p>\<^sub>0[r0, ?u1] \<cdot> ?RHS"
using RHS 2 C.prj_tuple [of r0 ?u1] by simp
finally show ?thesis by simp
qed
show "ru.prj\<^sub>1 \<cdot> ?LHS = ru.prj\<^sub>1 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>1 \<cdot> ?LHS =
(ru.prj\<^sub>1 \<cdot> r\<theta>'.chine) \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = \<p>\<^sub>1[r0, fw'.leg1] \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using Chn_r\<theta>' Chn_r\<theta>'_eq fw'
\<open>C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1']
(\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1'])\<close>
by simp
also have "... = (rfw'.Prj\<^sub>1 \<cdot> rfw'.chine_assoc) \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using ide_f ide_w' hseq_rf hseq_char \<alpha>_ide fw'.leg1_composite C.comp_assoc
by auto
also have "... = (rfw'.Prj\<^sub>1\<^sub>1 \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
using rfw'.prj_chine_assoc C.comp_assoc by simp
also have "... = ((k1 \<cdot> \<p>\<^sub>1[k0, ?w1']) \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
using C.comp_cod_arr by simp
also have "... = (k1 \<cdot> \<p>\<^sub>1[k0, ?w1'] \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = (k1 \<cdot> \<rho>.chine \<cdot> ?p1') \<cdot> \<beta>.chine"
using Chn_\<rho>w'_eq Dom_\<rho>_0 Cod_\<rho>_0
\<open>C.commutative_square k0 ?w1' (\<rho>.chine \<cdot> ?p1') ?p0'\<close>
by simp
also have "... = (k1 \<cdot> \<rho>.chine) \<cdot> ?p1' \<cdot> \<beta>.chine"
using C.comp_assoc by metis
also have "... = (?R \<cdot> ?p1') \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = ?p1' \<cdot> \<beta>.chine"
using C.comp_cod_arr C.prj1_in_hom [of ?R ?w1'] cospan' by simp
also have "... = ru.prj\<^sub>1 \<cdot> ?RHS"
using RHS 2 by simp
finally show ?thesis by simp
qed
qed
qed
finally show ?thesis by simp
qed
text \<open>
Now we work on the left-hand side.
\<close>
have L: "Chn (T.composite_cell w \<theta>) = \<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
proof -
have "Chn (T.composite_cell w \<theta>) = r\<theta>.chine \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using Chn_vcomp arfw C.comp_assoc by auto
moreover have "... = \<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
proof -
let ?LHS = "r\<theta>.chine \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
let ?RHS = "\<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
have 2: "C.commutative_square r0 ?u1 ?p1 \<theta>.chine" by fact
have LHS: "\<guillemotleft>?LHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r0 \<down>\<down> ?u1\<guillemotright>"
using Chn_r\<theta> Chn_\<rho>w rfw.chine_assoc_in_hom
by (metis (no_types, lifting) "8" Chn_in_hom Dom_\<rho>_0
arrow_of_spans_data.simps(2) calculation gw.chine_composite
r\<theta>_cod_apex_eq ru.chine_composite)
have RHS: "\<guillemotleft>?RHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r0 \<down>\<down> ?u1\<guillemotright>"
using 2 C.tuple_in_hom [of r0 ?u1 "?p1" \<theta>.chine] cospan r\<theta>_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\<theta>_cod_apex_eq by auto
show "C.seq ru.prj\<^sub>1 ?RHS"
using RHS r\<theta>_cod_apex_eq by auto
show "ru.prj\<^sub>0 \<cdot> ?LHS = ru.prj\<^sub>0 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>0 \<cdot> ?LHS = (ru.prj\<^sub>0 \<cdot> r\<theta>.chine) \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, f.leg1 \<cdot> fw.prj\<^sub>1]) \<cdot>
Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using Chn_r\<theta>_eq Dom_\<theta>_1 Cod_\<theta>_1 fw.leg1_composite
\<open>C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1]
(\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1])\<close>
by simp
also have "... = \<theta>.chine \<cdot> (\<p>\<^sub>0[r0, r0 \<cdot> ?p1] \<cdot> Chn \<a>[r, f, w]) \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = \<theta>.chine \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine"
proof -
have "Chn \<a>[r, f, w] = rfw.chine_assoc"
using ide_f ide_w hseq_rf hseq_char \<alpha>_ide by auto
moreover have "\<p>\<^sub>0[r0, r0 \<cdot> ?p1] \<cdot> rfw.chine_assoc =
\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>"
using rfw.chine_assoc_def
\<open>C.commutative_square r0 (r0 \<cdot> ?p1) rfw.Prj\<^sub>1\<^sub>1
\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>\<close>
by simp
ultimately show ?thesis by simp
qed
also have "... = \<theta>.chine \<cdot> (?R \<down>\<down> ?w1)"
proof -
have "\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine = ?R \<down>\<down> ?w1"
proof (intro C.prj_joint_monic
[of ?R ?w1 "\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine"
"?R \<down>\<down> ?w1"])
show "C.cospan ?R ?w1" by fact
show "C.seq ?p1 (\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine)"
proof -
have "C.seq rfw.Prj\<^sub>0\<^sub>1 \<rho>w.chine"
by (meson C.seqI' Chn_in_hom \<rho>w rfw.prj_in_hom(2)
\<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close>)
thus ?thesis
using \<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close>
by (metis (no_types) C.comp_assoc C.prj_tuple(2))
qed
show "C.seq ?p1 (?R \<down>\<down> ?w1)"
using gw.dom.apex_def gw.leg0_composite gw.prj_in_hom by auto
show "?p0 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
?p0 \<cdot> (?R \<down>\<down> ?w1)"
proof -
have "?p0 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
(?p0 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = rfw.Prj\<^sub>0 \<cdot> \<rho>w.chine"
using \<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close> by simp
also have "... = \<p>\<^sub>0[k0, ?w1] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using Chn_\<rho>w_eq C.comp_cod_arr by simp
also have "... = ?p0"
using \<open>C.commutative_square k0 ?w1 (\<rho>.chine \<cdot> ?p1) ?p0\<close>
C.prj_tuple(1)
by blast
also have "... = ?p0 \<cdot> (?R \<down>\<down> ?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 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
?p1 \<cdot> (?R \<down>\<down> ?w1)"
proof -
have "?p1 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
(?p1 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = rfw.Prj\<^sub>0\<^sub>1 \<cdot> \<rho>w.chine"
using \<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close> by simp
also have "... = (k0 \<cdot> \<p>\<^sub>1[k0, ?w1]) \<cdot> \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using Chn_\<rho>w_eq C.comp_cod_arr by simp
also have "... = k0 \<cdot> \<p>\<^sub>1[k0, ?w1] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using C.comp_assoc by simp
also have "... = k0 \<cdot> \<rho>.chine \<cdot> ?p1"
using \<open>C.commutative_square k0 ?w1 (\<rho>.chine \<cdot> ?p1) ?p0\<close> by simp
also have "... = (k0 \<cdot> \<rho>.chine) \<cdot> ?p1"
using C.comp_assoc by metis
also have "... = ?p1 \<cdot> (?R \<down>\<down> ?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 "... = \<theta>.chine"
using C.comp_arr_dom \<theta>.chine_in_hom gw.chine_eq_apex gw.chine_is_identity
Dom_\<theta>_0 Cod_\<theta>_0 Dom_\<theta>.apex_def Cod_\<theta>.apex_def
by (metis Dom_g \<theta>.chine_simps(1) \<theta>.chine_simps(2) gw.chine_composite
gw.dom.apex_def gw.leg0_composite span_data.select_convs(1))
also have "... = ru.prj\<^sub>0 \<cdot> ?RHS"
using 2 by simp
finally show ?thesis by blast
qed
show "ru.prj\<^sub>1 \<cdot> ?LHS = ru.prj\<^sub>1 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>1 \<cdot> ?LHS = (ru.prj\<^sub>1 \<cdot> r\<theta>.chine) \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = (r.chine \<cdot> \<p>\<^sub>1[r0, r0 \<cdot> ?p1]) \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
proof -
have "r\<theta>.chine \<noteq> C.null \<Longrightarrow>
\<p>\<^sub>1[r.cod.leg0, Cod_\<theta>.leg1] \<cdot> r\<theta>.chine =
r.chine \<cdot> \<p>\<^sub>1[r0, Dom_\<theta>.leg1]"
by (metis (lifting) C.prj_tuple(2) C.tuple_is_extensional r.cod_simps(2)
r\<theta>.chine_composite)
thus ?thesis
using Cod_\<theta>_1 Dom_\<theta>_1 r\<theta>.chine_simps(1) fw by fastforce
qed
also have "... = r.chine \<cdot> (rfw.Prj\<^sub>1 \<cdot> Chn \<a>[r, f, w]) \<cdot> \<rho>w.chine"
using C.comp_assoc fw.leg1_composite by simp
also have "... = r.chine \<cdot> rfw.Prj\<^sub>1\<^sub>1 \<cdot> \<rho>w.chine"
using ide_f ide_w hseq_rf hseq_char \<alpha>_ide
rfw.prj_chine_assoc(1)
by auto
also have "... = r.chine \<cdot> k1 \<cdot> \<p>\<^sub>1[k0, ?w1] \<cdot> \<rho>w.chine"
using C.comp_cod_arr C.comp_assoc by simp
also have "... = r.chine \<cdot> k1 \<cdot> \<rho>.chine \<cdot> \<p>\<^sub>1[Dom_\<rho>.leg0, ?w1]"
using Chn_\<rho>w_eq
\<open>C.commutative_square k0 ?w1
(\<rho>.chine \<cdot> \<p>\<^sub>1[ra, w.leg1]) \<p>\<^sub>0[ra, w.leg1]\<close>
by auto
also have "... = r.chine \<cdot> (k1 \<cdot> \<rho>.chine) \<cdot> ?p1"
using C.comp_assoc Dom_\<rho>_0 by metis
also have "... = r.chine \<cdot> ra \<cdot> ?p1"
by simp
also have "... = r.chine \<cdot> ?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 \<cdot> ?RHS"
using 2 by simp
finally show ?thesis by simp
qed
qed
qed
ultimately show ?thesis
by simp
qed
text \<open>
This is the main point: the equation E boils down to the following:
\[
\<open>?p1' \<cdot> \<beta>.chine = ?p1 \<and> \<theta>'.chine \<cdot> \<beta>.chine = \<theta>.chine\<close>
\]
The first equation gets us close to what we need, but we still need
\<open>?p1 \<cdot> C.inv ?p0 = ?w1\<close>, which follows from the fact that ?p0 is the
pullback of ?R.
\<close>
have *: "\<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle> = \<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
using L R E by simp
have **: "?p1' \<cdot> \<beta>.chine = ?p1"
by (metis "*" C.in_homE C.not_arr_null C.prj_tuple(2) C.tuple_in_hom
C.tuple_is_extensional
\<open>C.commutative_square r0 u.leg1
(\<p>\<^sub>1[ra, w'.leg1] \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)\<close>)
have ***: "\<theta>'.chine \<cdot> \<beta>.chine = \<theta>.chine"
by (metis "*" C.prj_tuple(1) \<open>C.commutative_square r0 ?u1
(?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)\<close>
\<open>C.commutative_square r0 ?u1 ?p1 \<theta>.chine\<close>)
text \<open>
CKS say to take \<open>\<gamma> = \<beta>\<close>, but obviously this cannot work as
literally described, because \<open>\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>\<close>, whereas we must have
\<open>\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>\<close>. Instead, we have to define \<open>\<gamma>\<close> by transporting \<open>\<beta>\<close> along the
projections from \<open>?R \<down>\<down> ?w1\<close> to \<open>?W\<close> and \<open>?R \<down>\<down> ?w1'\<close> to \<open>?W'\<close>.
These are isomorphisms by virtue of their being pullbacks of identities,
but they are not themselves necessarily identities.
Specifically, we take \<open>Chn \<gamma> = ?p0' \<cdot> Chn \<beta> \<cdot> C.inv ?p0\<close>.
\<close>
let ?\<gamma> = "\<lparr>Chn = ?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0, Dom = Dom w, Cod = Cod w'\<rparr>"
interpret Dom_\<gamma>: span_in_category C \<open>Dom ?\<gamma>\<close>
using w.dom.span_in_category_axioms by simp
interpret Cod_\<gamma>: span_in_category C \<open>Cod ?\<gamma>\<close>
using w'.cod.span_in_category_axioms by simp
text \<open>
It has to be shown that \<open>\<gamma>\<close> is an arrow of spans.
\<close>
interpret \<gamma>: arrow_of_spans C ?\<gamma>
proof
show "\<guillemotleft>Chn ?\<gamma> : Dom_\<gamma>.apex \<rightarrow>\<^sub>C Cod_\<gamma>.apex\<guillemotright>"
proof -
have "\<guillemotleft>Chn \<beta>: gw.apex \<rightarrow>\<^sub>C gw'.apex\<guillemotright>"
using Chn_in_hom \<beta> gw'.chine_eq_apex gw.chine_eq_apex by force
moreover have "\<guillemotleft>?p0' : gw'.apex \<rightarrow>\<^sub>C w'.apex\<guillemotright>"
using cospan' hseq_gw' hseq_char hcomp_def gw'.dom.apex_def w'.dom.apex_def
by auto
moreover have "\<guillemotleft>C.inv ?p0 : w.apex \<rightarrow>\<^sub>C gw.apex\<guillemotright>"
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_\<gamma>.apex_def Cod_\<gamma>.apex_def by auto
qed
text \<open>
The commutativity property for the ``input leg'' follows directly from that
for \<open>\<beta>\<close>.
\<close>
show "Cod_\<gamma>.leg0 \<cdot> Chn ?\<gamma> = Dom_\<gamma>.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_\<beta>.leg_simps(1) Dom_\<beta>_eq \<beta>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 \<open>
The commutativity property for the ``output leg'' is a bit more subtle.
\<close>
show "Cod_\<gamma>.leg1 \<cdot> Chn ?\<gamma> = Dom_\<gamma>.leg1"
proof -
have "Cod_\<gamma>.leg1 \<cdot> Chn ?\<gamma> = ((?w1' \<cdot> ?p0') \<cdot> \<beta>.chine) \<cdot> C.inv ?p0"
using C.comp_assoc by simp
also have "... = ((?R \<cdot> ?p1') \<cdot> Chn \<beta>) \<cdot> C.inv ?p0"
using cospan' C.pullback_commutes [of ?R ?w1'] by auto
also have "... = (?p1' \<cdot> \<beta>.chine) \<cdot> C.inv ?p0"
using cospan' C.comp_cod_arr by simp
also have "... = ?p1 \<cdot> C.inv ?p0"
using ** by simp
also have "... = ?w1"
text \<open>
Sledgehammer found this at a time when I was still struggling to
understand what was going on.
\<close>
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_\<gamma>.leg1" by auto
finally show ?thesis by simp
qed
qed
text \<open>
What remains to be shown is that \<open>\<gamma>\<close> is unique with the properties asserted
by \<open>T2\<close>; \emph{i.e.} \<open>\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>)\<close>.
CKS' assertion that the equation \<open>(r\<theta>)(\<rho>w) = (r\<theta>')(\<rho>w')\<beta>\<close> gives \<open>w\<^sub>1 = w\<^sub>1'\<close>
does not really seem to be true. The reason \<open>\<gamma>\<close> is unique is because it is
obtained by transporting \<open>\<beta>\<close> along isomorphisms.
\<close>
have \<gamma>: "\<guillemotleft>?\<gamma> : w \<Rightarrow> w'\<guillemotright>"
using \<gamma>.arrow_of_spans_axioms arr_char dom_char cod_char by auto
have hseq_f\<gamma>: "hseq f ?\<gamma>"
using \<gamma> src_def trg_def arrI fw.composable rf.are_arrows(2) by auto
have hseq_g\<gamma>: "hseq g ?\<gamma>"
using \<gamma> src_def trg_def fw.composable gw.are_arrows(1) src_f by auto
interpret f\<gamma>: two_composable_arrows_of_spans C prj0 prj1 f ?\<gamma>
using hseq_f\<gamma> hseq_char by (unfold_locales, simp)
interpret f\<gamma>: arrow_of_spans C \<open>f \<star> ?\<gamma>\<close>
using f\<gamma>.composite_is_arrow arr_char by simp
interpret g\<gamma>: two_composable_arrows_of_spans C prj0 prj1 g ?\<gamma>
using hseq_g\<gamma> hseq_char by (unfold_locales, simp)
interpret g\<gamma>: arrow_of_spans C \<open>g \<star> ?\<gamma>\<close>
using g\<gamma>.composite_is_arrow arr_char by simp
have Chn_g\<gamma>: "Chn (g \<star> ?\<gamma>) = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> ?p0' \<cdot> \<beta>.chine\<rangle>"
proof -
have "Chn (g \<star> ?\<gamma>) = \<langle>?R \<cdot> ?p1 \<lbrakk>?R, ?w1'\<rbrakk> (?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0) \<cdot> ?p0\<rangle>"
using g\<gamma>.chine_composite by simp
also have "... = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> (?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0) \<cdot> ?p0\<rangle>"
using C.comp_cod_arr cospan by simp
also have "... = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> ?p0' \<cdot> \<beta>.chine\<rangle>"
proof -
have "(?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0) \<cdot> ?p0 = ?p0' \<cdot> \<beta>.chine"
using C.comp_assoc C.iso_pullback_ide [of ?R ?w1] C.comp_inv_arr
C.comp_arr_dom Chn_\<beta>
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_\<beta>_eq: "\<beta>.chine = Chn (g \<star> ?\<gamma>)"
by (metis "**" C.span_prj C.tuple_prj Chn_g\<gamma> cospan cospan')
have \<beta>_eq_g\<gamma>: "\<beta> = g \<star> ?\<gamma>"
proof (intro arr_eqI)
show "par \<beta> (g \<star> ?\<gamma>)"
proof -
have "\<guillemotleft>g \<star> ?\<gamma> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
using ide_g \<gamma> T.leg1_simps(3)
by (intro hcomp_in_vhom, auto)
thus ?thesis
using \<beta> by (elim in_homE, auto)
qed
show "\<beta>.chine = Chn (g \<star> ?\<gamma>)"
using Chn_\<beta>_eq by simp
qed
moreover have "\<theta> = \<theta>' \<bullet> (f \<star> ?\<gamma>)"
proof (intro arr_eqI)
have f\<gamma>: "\<guillemotleft>f \<star> ?\<gamma> : f \<star> w \<Rightarrow> f \<star> w'\<guillemotright>"
using \<gamma> ide_f by auto
show par: "par \<theta> (\<theta>' \<bullet> (f \<star> ?\<gamma>))"
using \<theta> \<theta>' f\<gamma> by (elim in_homE, auto)
show "\<theta>.chine = Chn (\<theta>' \<bullet> (f \<star> ?\<gamma>))"
using par "***" Chn_vcomp calculation f\<gamma>.chine_composite g\<gamma>.chine_composite
by auto
qed
ultimately show 2: "\<guillemotleft>?\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> ?\<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> ?\<gamma>)"
using \<gamma> by simp
show "\<And>\<gamma>'. \<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma>' \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>') \<Longrightarrow> \<gamma>' = ?\<gamma>"
proof -
fix \<gamma>'
assume 1: "\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma>' \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>')"
interpret \<gamma>': arrow_of_spans C \<gamma>'
using 1 arr_char by auto
have hseq_g\<gamma>': \<open>hseq g \<gamma>'\<close>
using 1 \<beta> by auto
interpret g\<gamma>': two_composable_arrows_of_spans C prj0 prj1 g \<gamma>'
using hseq_g\<gamma>' hseq_char by unfold_locales auto
interpret g\<gamma>': arrow_of_spans C \<open>g \<star> \<gamma>'\<close>
using g\<gamma>'.composite_is_arrow arr_char by simp
show "\<gamma>' = ?\<gamma>"
proof (intro arr_eqI)
show par: "par \<gamma>' ?\<gamma>"
using 1 \<gamma> by fastforce
show "\<gamma>'.chine = \<gamma>.chine"
proof -
have "C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>'.chine \<cdot> ?p0)"
proof
show "C.cospan ?R ?w1'" by fact
show 3: "C.span (g.chine \<cdot> ?p1) (\<gamma>'.chine \<cdot> ?p0)"
proof (intro conjI)
show "C.seq g.chine ?p1"
using cospan by auto
show "C.seq \<gamma>'.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 \<cdot> ?p1) = C.dom (\<gamma>'.chine \<cdot> ?p0)"
using g.chine_eq_apex cospan by simp
qed
show "C.dom ra = C.cod (g.chine \<cdot> ?p1)"
using cospan by auto
show "?R \<cdot> g.chine \<cdot> ?p1 = ?w1' \<cdot> \<gamma>'.chine \<cdot> ?p0"
proof -
have "?w1' \<cdot> \<gamma>'.chine \<cdot> ?p0 = (?w1' \<cdot> \<gamma>'.chine) \<cdot> ?p0"
using C.comp_assoc by simp
moreover have "... = ?w1 \<cdot> ?p0"
using 1 \<gamma>'.leg1_commutes dom_char cod_char by auto
also have "... = ?R \<cdot> ?p1"
using cospan C.pullback_commutes [of ra ?w1] by auto
also have "... = ?R \<cdot> g.chine \<cdot> ?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 \<cdot> ?p1) (\<gamma>.chine \<cdot> ?p0)"
proof
show "C.cospan ?R ?w1'" by fact
show 3: "C.span (g.chine \<cdot> ?p1) (\<gamma>.chine \<cdot> ?p0)"
using cospan \<gamma>.chine_in_hom by auto
show "C.dom ?R = C.cod (g.chine \<cdot> ?p1)"
using cospan by auto
show "?R \<cdot> g.chine \<cdot> ?p1 = ?w1' \<cdot> \<gamma>.chine \<cdot> ?p0"
proof -
have "?w1' \<cdot> \<gamma>.chine \<cdot> ?p0 = (?w1' \<cdot> \<gamma>.chine) \<cdot> ?p0"
using C.comp_assoc by simp
moreover have "... = ?w1 \<cdot> ?p0"
using 1 \<gamma>.leg1_commutes dom_char cod_char by auto
also have "... = ?R \<cdot> ?p1"
using cospan C.pullback_commutes [of ra ?w1] by auto
also have "... = ?R \<cdot> g.chine \<cdot> ?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 "\<gamma>'.chine \<cdot> ?p0 = \<gamma>.chine \<cdot> ?p0"
proof -
have "\<gamma>'.chine \<cdot> ?p0 = ?p0' \<cdot> g\<gamma>'.chine"
using 1 dom_char cod_char g\<gamma>'.chine_composite
\<open>C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>'.chine \<cdot> ?p0)\<close>
by auto
also have "... = ?p0' \<cdot> \<beta>.chine"
using 1 by simp
also have "... = ?p0' \<cdot> g\<gamma>.chine"
using Chn_\<beta>_eq by simp
also have "... = \<gamma>.chine \<cdot> ?p0"
using g\<gamma>.chine_composite
\<open>C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>.chine \<cdot> ?p0)\<close>
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" \<gamma>'.chine \<gamma>.chine] cospan \<gamma>.chine_in_hom
\<gamma>'.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 "\<cong>" 50)
text \<open>
\<open>Span(C)\<close> is a bicategory of spans.
\<close>
lemma is_bicategory_of_spans:
shows "bicategory_of_spans vcomp hcomp assoc unit src trg"
proof
text \<open>
Every 1-cell \<open>r\<close> is isomorphic to the composition of a map and the right adjoint
of a map. The proof is to obtain a tabulation of \<open>r\<close> as a span of maps \<open>(f, g)\<close>
and then observe that \<open>r\<close> is isomorphic to \<open>g \<star> f\<^sup>*\<close>.
\<close>
show "\<And>r. ide r \<Longrightarrow> \<exists>f g. is_left_adjoint f \<and> is_left_adjoint g \<and> r \<cong> g \<star> 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 \<rho>: "tabulation (\<bullet>) (\<star>) assoc unit src trg r r.\<rho> r.f r.g \<and>
is_left_adjoint r.f \<and> is_left_adjoint r.g"
using r r.has_tabulation by blast
interpret \<rho>: tabulation vcomp hcomp assoc unit src trg r r.\<rho> r.f r.g
using \<rho> by fast
have 1: "r \<cong> r.g \<star> r.f\<^sup>*"
using \<rho> \<rho>.yields_isomorphic_representation' \<rho>.T0.is_map
left_adjoint_extends_to_adjoint_pair
isomorphic_def [of "r.g \<star> r.f\<^sup>*" r] isomorphic_symmetric
by auto
thus "\<exists>f g. is_left_adjoint f \<and> is_left_adjoint g \<and> r \<cong> g \<star> f\<^sup>*"
using \<rho> by blast
qed
text \<open>
Every span of maps extends to a tabulation.
\<close>
show "\<And>f g. \<lbrakk> is_left_adjoint f; is_left_adjoint g; src f = src g \<rbrakk> \<Longrightarrow>
\<exists>r \<rho>. tabulation (\<bullet>) (\<star>) assoc unit src trg r \<rho> f g"
proof -
text \<open>
The proof idea is as follows: Let maps \<open>f = (f\<^sub>1, f\<^sub>0)\<close> and \<open>g = (g\<^sub>1, g\<^sub>0)\<close> be given.
Let \<open>f' = (f\<^sub>1 \<cdot> C.inv f\<^sub>0, C.cod f\<^sub>0)\<close> and \<open>g' = (g\<^sub>1 \<cdot> C.inv g\<^sub>0, C.cod g\<^sub>0)\<close>;
then \<open>f'\<close> and \<open>g'\<close> are maps isomorphic to \<open>f\<close> and \<open>g\<close>, respectively.
By a previous result, \<open>f'\<close> and \<open>g'\<close> extend to a tabulation \<open>(f', \<tau>, g')\<close> of
\<open>r = (f\<^sub>1 \<cdot> C.inv f\<^sub>0, g\<^sub>1 \<cdot> C.inv g\<^sub>0)\<close>.
Compose with isomorphisms \<open>\<guillemotleft>\<phi> : f' \<Rightarrow> f\<guillemotright>\<close> and \<open>\<guillemotleft>\<psi> : g \<Rightarrow> g'\<guillemotright>\<close> to obtain
\<open>(f, (r \<star> \<phi>) \<cdot> \<tau> \<cdot> \<psi>, g)\<close> and show it must also be a tabulation.
\<close>
fix f g
assume f: "is_left_adjoint f"
assume g: "is_left_adjoint g"
assume fg: "src f = src g"
show "\<exists>r \<rho>. tabulation (\<bullet>) (\<star>) assoc unit src trg r \<rho> f g"
proof -
text \<open>We have to unpack the hypotheses to get information about f and g.\<close>
obtain f\<^sub>a \<eta>\<^sub>f \<epsilon>\<^sub>f
where ff\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \<eta>\<^sub>f \<epsilon>\<^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 \<eta>\<^sub>f \<epsilon>\<^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 \<eta>\<^sub>g \<epsilon>\<^sub>g
where G: "adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \<eta>\<^sub>g \<epsilon>\<^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 \<eta>\<^sub>g \<epsilon>\<^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 \<cdot> C.inv f.leg0)"
have f': "ide ?f'"
proof -
have "C.span (C.cod f.leg0) (f.leg1 \<cdot> 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 \<cdot> C.inv g.leg0)"
have g': "ide ?g'"
proof -
have "C.span (C.cod g.leg0) (g.leg1 \<cdot> 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 \<rho>: tabulation \<open>(\<bullet>)\<close> \<open>(\<star>)\<close> assoc unit src trg ?r r.\<rho> r.f r.g
using r.has_tabulation by simp
have \<rho>_eq: "r.\<rho> = \<lparr>Chn = \<langle>C.cod f.leg0 \<lbrakk>f'.leg1, f'.leg1\<rbrakk> C.cod f.leg0\<rangle>,
Dom = \<lparr>Leg0 = C.cod f.leg0, Leg1 = g'.leg1\<rparr>,
Cod = \<lparr>Leg0 = \<p>\<^sub>0[f'.leg1, f'.leg1],
Leg1 = g'.leg1 \<cdot> \<p>\<^sub>1[f'.leg1, f'.leg1]\<rparr>\<rparr>"
using \<open>r.f = ?f'\<close> by auto
text \<open>Obtain the isomorphism from \<open>f'\<close> to \<open>f\<close>.\<close>
let ?\<phi> = "\<lparr>Chn = C.inv f.leg0, Dom = Dom ?f', Cod = Dom f\<rparr>"
interpret Dom_\<phi>: span_in_category C
\<open>Dom \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr>\<close>
using f'.dom.span_in_category_axioms by simp
interpret Cod_\<phi>: span_in_category C
\<open>Cod \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr>\<close>
using f.dom.span_in_category_axioms by simp
interpret \<phi>: arrow_of_spans C ?\<phi>
proof
show "\<guillemotleft>Chn \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr> : Dom_\<phi>.apex \<rightarrow>\<^sub>C Cod_\<phi>.apex\<guillemotright>"
using f f.dom.apex_def f'.dom.apex_def is_left_adjoint_char by auto
show "Cod_\<phi>.leg0 \<cdot> Chn \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr> =
Dom_\<phi>.leg0"
using f f.dom.apex_def is_left_adjoint_char C.comp_arr_inv C.inv_is_inverse
by simp
show "Cod_\<phi>.leg1 \<cdot> Chn \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr> =
Dom_\<phi>.leg1"
by simp
qed
have \<phi>: "\<guillemotleft>?\<phi> : ?f' \<Rightarrow> f\<guillemotright> \<and> iso ?\<phi>"
using f is_left_adjoint_char iso_char arr_char dom_char cod_char
\<phi>.arrow_of_spans_axioms f'.dom.apex_def f.dom.apex_def
by auto
text \<open>
Obtain the isomorphism from \<open>g\<close> to \<open>g'\<close>.
Recall: \<open>g' = mkIde (C.cod g.leg0) (g.dom.leg1 \<cdot> C.inv g.leg0)\<close>.
The isomorphism is given by \<open>g.leg0\<close>.
\<close>
let ?\<psi> = "\<lparr>Chn = g.leg0, Dom = Dom g, Cod = Dom ?g'\<rparr>"
interpret Dom_\<psi>: span_in_category C
\<open>Dom \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr>\<close>
using g.dom.span_in_category_axioms by simp
interpret Cod_\<psi>: span_in_category C
\<open>Cod \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr>\<close>
using g'.dom.span_in_category_axioms by simp
interpret \<psi>: arrow_of_spans C ?\<psi>
proof
show "\<guillemotleft>Chn \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr> :
Dom_\<psi>.apex \<rightarrow>\<^sub>C Cod_\<psi>.apex\<guillemotright>"
using g g.dom.apex_def g'.dom.apex_def is_left_adjoint_char by auto
show "Cod_\<psi>.leg0 \<cdot> Chn \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr> =
Dom_\<psi>.leg0"
using C.comp_cod_arr by simp
show "Cod_\<psi>.leg1 \<cdot> Chn \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr> =
Dom_\<psi>.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 \<psi>: "\<guillemotleft>?\<psi> : g \<Rightarrow> ?g'\<guillemotright> \<and> iso ?\<psi>"
using g is_left_adjoint_char iso_char arr_char dom_char cod_char
\<psi>.arrow_of_spans_axioms g.dom.apex_def g'.dom.apex_def
by auto
have \<rho>\<psi>: "tabulation (\<bullet>) (\<star>) assoc unit src trg ?r (r.\<rho> \<bullet> ?\<psi>) r.f g"
using \<psi> \<open>r.g = ?g'\<close> r.has_tabulation \<rho>.preserved_by_output_iso by simp
interpret \<tau>\<psi>: tabulation vcomp hcomp assoc unit src trg ?r \<open>r.\<rho> \<bullet> ?\<psi>\<close> r.f g
using \<rho>\<psi> by auto
have "tabulation (\<bullet>) (\<star>) assoc unit src trg ?r ((?r \<star> ?\<phi>) \<bullet> r.\<rho> \<bullet> ?\<psi>) f g"
using \<phi> \<open>r.f = ?f'\<close> \<tau>\<psi>.preserved_by_input_iso [of ?\<phi> f] by argo
thus ?thesis by auto
qed
qed
text \<open>The sub-bicategory of maps is locally essentially discrete.\<close>
show "\<And>f f' \<mu> \<mu>'. \<lbrakk> is_left_adjoint f; is_left_adjoint f'; \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright>; \<guillemotleft>\<mu>' : f \<Rightarrow> f'\<guillemotright> \<rbrakk>
\<Longrightarrow> iso \<mu> \<and> iso \<mu>' \<and> \<mu> = \<mu>'"
proof -
fix f f' \<mu> \<mu>'
assume f: "is_left_adjoint f" and f': "is_left_adjoint f'"
assume \<mu>: "\<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright>" and \<mu>': "\<guillemotleft>\<mu>' : f \<Rightarrow> f'\<guillemotright>"
obtain f\<^sub>a \<eta> \<epsilon>
where f\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \<eta> \<epsilon>"
using f adjoint_pair_def by auto
obtain f'\<^sub>a \<eta>' \<epsilon>'
where f'\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \<eta>' \<epsilon>'"
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 \<eta> \<epsilon>
using f\<^sub>a by simp
interpret f'\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \<eta>' \<epsilon>'
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 \<mu>: arrow_of_spans C \<mu> using \<mu> arr_char by auto
interpret \<mu>': arrow_of_spans C \<mu>' using \<mu>' arr_char by auto
have 1: "C.iso f.leg0 \<and> C.iso f'.leg0"
using f f' is_left_adjoint_char by simp
have 2: "\<mu>.chine = C.inv f'.leg0 \<cdot> f.leg0"
using \<mu> 1 dom_char cod_char \<mu>.leg0_commutes C.invert_side_of_triangle by auto
moreover have "\<mu>'.chine = C.inv f'.leg0 \<cdot> f.leg0"
using \<mu>' 1 dom_char cod_char \<mu>'.leg0_commutes C.invert_side_of_triangle by auto
ultimately have 3: "\<mu>.chine = \<mu>'.chine" by simp
have "iso \<mu>"
using 1 2 \<mu> C.isos_compose dom_char cod_char iso_char arr_char by auto
hence "iso \<mu>'"
using 3 iso_char arr_char \<mu>'.arrow_of_spans_axioms by simp
moreover have "\<mu> = \<mu>'"
using 3 \<mu> \<mu>' dom_char cod_char by fastforce
ultimately show "iso \<mu> \<and> iso \<mu>' \<and> \<mu> = \<mu>'"
by simp
qed
qed
text \<open>
We can now prove the easier half of the main result (CKS Theorem 4):
If \<open>B\<close> is biequivalent to \<open>Span(C)\<close>, where \<open>C\<close> is a category with pullbacks,
then \<open>B\<close> 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.)
\<close>
theorem equivalent_implies_bicategory_of_spans:
assumes "equivalent_bicategories vcomp hcomp assoc unit src trg V\<^sub>1 H\<^sub>1 \<a>\<^sub>1 \<i>\<^sub>1 src\<^sub>1 trg\<^sub>1"
shows "bicategory_of_spans V\<^sub>1 H\<^sub>1 \<a>\<^sub>1 \<i>\<^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 \<open>
We now develop consequences of the axioms for a bicategory of spans, in preparation for
proving the other half of the main result.
\<close>
context bicategory_of_spans
begin
notation isomorphic (infix "\<cong>" 50)
text \<open>
The following is a convenience version of \<open>BS2\<close> that gives us what we generally want:
given specified \<open>f, g\<close> obtain \<open>\<rho>\<close> that makes \<open>(f, \<rho>, g)\<close> a tabulation of \<open>g \<star> f\<^sup>*\<close>,
not a tabulation of some \<open>r\<close> isomorphic to \<open>g \<star> f\<^sup>*\<close>.
\<close>
lemma BS2':
assumes "is_left_adjoint f" and "is_left_adjoint g" and "src f = src g"
and "isomorphic (g \<star> f\<^sup>*) r"
shows "\<exists>\<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"
proof -
have 1: "is_left_adjoint f \<and> is_left_adjoint g \<and> g \<star> f\<^sup>* \<cong> r"
using assms BS1 by simp
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"
using 1 isomorphic_def by blast
obtain r' \<rho>' where \<rho>': "tabulation V H \<a> \<i> src trg r' \<rho>' f g"
using assms 1 BS2 by blast
interpret \<rho>': tabulation V H \<a> \<i> src trg r' \<rho>' f g
using \<rho>' by simp
let ?\<psi> = "\<rho>'.T0.trnr\<^sub>\<epsilon> r' \<rho>'"
have \<psi>: "\<guillemotleft>?\<psi> : g \<star> f\<^sup>* \<Rightarrow> r'\<guillemotright> \<and> iso ?\<psi>"
using \<rho>'.yields_isomorphic_representation by blast
have "\<guillemotleft>\<phi> \<cdot> inv ?\<psi> : r' \<Rightarrow> r\<guillemotright> \<and> iso (\<phi> \<cdot> inv ?\<psi>)"
using \<phi> \<psi> isos_compose by blast
hence 3: "tabulation V H \<a> \<i> src trg r ((\<phi> \<cdot> inv ?\<psi> \<star> f) \<cdot> \<rho>') f g"
using \<rho>'.is_preserved_by_base_iso by blast
hence "\<exists>\<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"
by blast
thus ?thesis
using someI_ex [of "\<lambda>\<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"] by simp
qed
text \<open>
The following observation is made by CKS near the beginning of the proof of Theorem 4:
If \<open>w\<close> is an arbitrary 1-cell, and \<open>g\<close> and \<open>g \<star> w\<close> are maps, then \<open>w\<close> is in fact a map.
It is applied frequently.
\<close>
lemma BS4:
assumes "is_left_adjoint g" and "ide w" and "is_left_adjoint (g \<star> w)"
shows "is_left_adjoint w"
proof -
text \<open>
CKS say: ``by (i) there are maps \<open>m, n\<close> with \<open>w \<cong> nm\<^sup>*\<close>, so, by (ii), we have two
tabulations \<open>(1, \<rho>, gw)\<close>, \<open>(m, \<sigma>, gn)\<close> of \<open>gw\<close>; since tabulations are unique
up to equivalence, \<open>m\<close> is invertible and \<open>w \<cong> nm\<^sup>*\<close> is a map.''
\<close>
have ex_\<rho>: "\<exists>\<rho>. tabulation V H \<a> \<i> src trg (g \<star> w) \<rho> (src w) (g \<star> w)"
proof -
have "(g \<star> w) \<star> src w \<cong> g \<star> w"
by (metis assms(3) iso_runit ideD(1) isomorphic_def left_adjoint_is_ide
runit_in_hom(2) src_hcomp)
moreover have "(g \<star> w) \<star> (src w)\<^sup>* \<cong> g \<star> w"
proof -
have "(g \<star> w) \<star> src (g \<star> w) \<cong> g \<star> w"
using calculation isomorphic_implies_ide(2) by auto
moreover have "src (g \<star> w) \<cong> (src w)\<^sup>*"
proof -
interpret src_w: map_in_bicategory V H \<a> \<i> src trg \<open>src w\<close>
using assms obj_is_self_adjoint by unfold_locales auto
interpret src_w: adjunction_in_bicategory V H \<a> \<i> src trg
\<open>src w\<close> \<open>(src w)\<^sup>*\<close> src_w.\<eta> src_w.\<epsilon>
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 \<cong> (src w)\<^sup>*"
using left_adjoint_determines_right_up_to_iso by simp
moreover have "src w = src (g \<star> w)"
using assms isomorphic_def hcomp_simps(1) left_adjoint_is_ide by simp
ultimately show ?thesis by simp
qed
moreover have "src (g \<star> w) = trg (src (g \<star> 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 \<star> w" "g \<star> w"]
by auto
qed
obtain \<rho> where \<rho>: "tabulation V H \<a> \<i> src trg (g \<star> w) \<rho> (src w) (g \<star> w)"
using ex_\<rho> by auto
obtain m n where mn: "is_left_adjoint m \<and> is_left_adjoint n \<and> isomorphic w (n \<star> m\<^sup>*)"
using assms BS1 [of w] by auto
have m\<^sub>a: "adjoint_pair m m\<^sup>* \<and> isomorphic w (n \<star> m\<^sup>*)"
using mn adjoint_pair_def left_adjoint_extends_to_adjoint_pair by blast
have ex_\<sigma>: "\<exists>\<sigma>. tabulation V H \<a> \<i> src trg (g \<star> w) \<sigma> m (g \<star> n)"
proof -
have "hseq n m\<^sup>*"
using mn isomorphic_implies_ide by auto
have "trg (n \<star> 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 \<open>hseq n m\<^sup>*\<close> adjoint_pair_antipar [of m "m\<^sup>*"] by fastforce
have "is_left_adjoint (g \<star> n)"
using assms mn left_adjoints_compose \<open>hseq g n\<close> by blast
moreover have "src m = src (g \<star> n)"
using assms mn \<open>hseq g n\<close> \<open>src m = src n\<close> by simp
moreover have "(g \<star> n) \<star> m\<^sup>* \<cong> g \<star> w"
proof -
have 1: "src g = trg (n \<star> m\<^sup>*)"
using assms \<open>trg (n \<star> m\<^sup>*) = trg w\<close> \<open>hseq g w\<close> by fastforce
hence "(g \<star> n) \<star> m\<^sup>* \<cong> g \<star> n \<star> m\<^sup>*"
using assms mn m\<^sub>a assoc_in_hom iso_assoc \<open>hseq g n\<close> \<open>hseq n m\<^sup>*\<close>
isomorphic_def left_adjoint_is_ide right_adjoint_is_ide
by (metis hseqE ideD(2) ideD(3))
also have "... \<cong> g \<star> 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 \<sigma> where \<sigma>: "tabulation V H \<a> \<i> src trg (g \<star> w) \<sigma> m (g \<star> n)"
using ex_\<sigma> by auto
interpret \<rho>: tabulation V H \<a> \<i> src trg \<open>g \<star> w\<close> \<rho> \<open>src w\<close> \<open>g \<star> w\<close>
using \<rho> by auto
interpret \<sigma>: tabulation V H \<a> \<i> src trg \<open>g \<star> w\<close> \<sigma> m \<open>g \<star> n\<close>
using \<sigma> by auto
text \<open>
As usual, the sketch given by CKS seems more suggestive than it is a precise recipe.
We can obtain an equivalence map \<open>\<guillemotleft>e : src w \<rightarrow> src m\<guillemotright>\<close> and \<open>\<theta>\<close> such that
\<open>\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright>\<close>.
We can also obtain an equivalence map \<open>\<guillemotleft>e' : src m \<rightarrow> src w\<guillemotright>\<close> and \<open>\<theta>'\<close> such that
\<open>\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright>\<close>. If \<open>\<theta>'\<close> can be taken to be an isomorphism; then we have
\<open>e' \<cong> src w \<star> e' \<cong> m\<close>. Since \<open>e'\<close> is an equivalence, this shows \<open>m\<close> is an equivalence,
hence its right adjoint \<open>m\<^sup>*\<close> is also an equivalence and therefore a map.
But \<open>w = n \<star> m\<^sub>a\<close>, so this shows that \<open>w\<close> is a map.
Now, we may assume without loss of generality that \<open>e\<close> and \<open>e'\<close> are part of an
adjoint equivalence.
We have \<open>\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright>\<close> and \<open>\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright>\<close>.
We may take the transpose of \<open>\<theta>\<close> to obtain \<open>\<guillemotleft>\<zeta> : m \<Rightarrow> src w \<star> e'\<guillemotright>\<close>;
then \<open>\<guillemotleft>\<theta>' \<cdot> \<zeta> : m \<Rightarrow> m\<guillemotright>\<close> and \<open>\<guillemotleft>\<zeta> \<cdot> \<theta>' : src w \<star> e' \<Rightarrow> src w \<star> e'\<guillemotright>\<close>.
Since \<open>m\<close> and \<open>src w \<star> e'\<close> are maps, by \<open>BS3\<close> it must be that \<open>\<zeta>\<close> and \<open>\<theta>'\<close> are inverses.
\<close>
text \<open>
{\bf Note:} CKS don't cite \<open>BS3\<close> here. I am not sure whether this result can be proved
without \<open>BS3\<close>. 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 \<open>BS3\<close> is not used in the proof below.
\<close>
have 2: "\<exists>e e' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'.
equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg e e' \<phi> \<psi> \<and>
\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<star> n \<Rightarrow> (g \<star> w) \<star> e'\<guillemotright> \<and> iso \<nu> \<and>
\<sigma> = \<rho>.composite_cell e' \<theta>' \<cdot> \<nu> \<and>
\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright> \<and> \<guillemotleft>\<nu>' : g \<star> w \<Rightarrow> (g \<star> n) \<star> e\<guillemotright> \<and> iso \<nu>' \<and>
\<rho> = ((g \<star> w) \<star> \<theta>) \<cdot> \<a>[g \<star> w, m, e] \<cdot> (\<sigma> \<star> e) \<cdot> \<nu>'"
using \<rho> \<sigma>.apex_unique_up_to_equivalence [of \<rho> "src w" "g \<star> w"] comp_assoc
by metis
obtain e e' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'
where *: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg e e' \<phi> \<psi> \<and>
\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<star> n \<Rightarrow> (g \<star> w) \<star> e'\<guillemotright> \<and> iso \<nu> \<and>
\<sigma> = \<rho>.composite_cell e' \<theta>' \<cdot> \<nu> \<and>
\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright> \<and> \<guillemotleft>\<nu>' : g \<star> w \<Rightarrow> (g \<star> n) \<star> e\<guillemotright> \<and> iso \<nu>' \<and>
\<rho> = \<sigma>.composite_cell e \<theta> \<cdot> \<nu>'"
using 2 comp_assoc by auto
interpret ee': equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg e e' \<phi> \<psi>
using * by simp
have equiv_e: "equivalence_map e"
using ee'.equivalence_in_bicategory_axioms equivalence_map_def by auto
obtain \<psi>' where \<psi>': "adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg e e' \<phi> \<psi>'"
using equivalence_refines_to_adjoint_equivalence [of e e' \<phi>]
ee'.unit_in_hom(2) ee'.unit_is_iso ee'.antipar equiv_e
by auto
interpret ee': adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg e e' \<phi> \<psi>'
using \<psi>' by simp
interpret e'e: adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg e' e \<open>inv \<psi>'\<close> \<open>inv \<phi>\<close>
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 \<theta>]
by (elim conjE in_homE) simp
have "hseq (src w) e'"
using * ide_dom [of \<theta>']
by (elim conjE in_homE) simp
have "e'e.trnr\<^sub>\<eta> m \<theta> \<in> hom m (src w \<star> e')"
proof -
have "src m = trg e"
using \<open>hseq m e\<close> by auto
moreover have "src (src w) = trg e'"
using \<open>hseq (src w) e'\<close> 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: "\<guillemotleft>e'e.trnr\<^sub>\<eta> m \<theta> : m \<Rightarrow> src w \<star> e'\<guillemotright>"
by simp
hence "\<guillemotleft>\<theta>' \<cdot> e'e.trnr\<^sub>\<eta> m \<theta> : m \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>e'e.trnr\<^sub>\<eta> m \<theta> \<cdot> \<theta>' : src w \<star> e' \<Rightarrow> src w \<star> e'\<guillemotright>"
using * by auto
moreover have "\<guillemotleft>m : m \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>src w \<star> e' : src w \<star> e' \<Rightarrow> src w \<star> e'\<guillemotright>"
using mn 3 ide_cod [of "e'e.trnr\<^sub>\<eta> m \<theta>"] left_adjoint_is_ide by fastforce
moreover have 4: "is_left_adjoint (src w \<star> 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 \<open>hseq (src w) e'\<close> by auto
qed
ultimately have "\<theta>' \<cdot> e'e.trnr\<^sub>\<eta> m \<theta> = m \<and> e'e.trnr\<^sub>\<eta> m \<theta> \<cdot> \<theta>' = src w \<star> e'"
using mn BS3 [of m m "\<theta>' \<cdot> e'e.trnr\<^sub>\<eta> m \<theta>" m]
BS3 [of "src w \<star> e'" "src w \<star> e'" "e'e.trnr\<^sub>\<eta> m \<theta> \<cdot> \<theta>'" "src w \<star> e'"]
by auto
hence "inverse_arrows \<theta>' (e'e.trnr\<^sub>\<eta> m \<theta>)"
using mn 4 left_adjoint_is_ide inverse_arrows_def by simp
hence 5: "iso \<theta>'"
by auto
have "equivalence_map (src w \<star> e')"
using assms obj_is_equivalence_map equiv_e' \<open>hseq (src w) e'\<close> 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 \<star> 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 "\<cong>" 50)
notation iso_class ("\<lbrakk>_\<rbrakk>")
text \<open>
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 \<open>BS1\<close> to choose a specific factorization of the form
\<open>r \<cong> tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*\<close> for each 1-cell \<open>r\<close>. This has to be done in such a way
that all elements of an isomorphism class are assigned the same factorization.
\<close>
abbreviation isomorphic_rep
where "isomorphic_rep r f g \<equiv> is_left_adjoint f \<and> is_left_adjoint g \<and> g \<star> f\<^sup>* \<cong> r"
definition tab\<^sub>0
where "tab\<^sub>0 r \<equiv> SOME f. \<exists>g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) f g"
definition tab\<^sub>1
where "tab\<^sub>1 r \<equiv> SOME g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) (tab\<^sub>0 r) g"
definition rep
where "rep r \<equiv> SOME \<phi>. \<guillemotleft>\<phi> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"
lemma rep_props:
assumes "ide r"
shows "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>" and "iso (rep r)"
and "r \<cong> iso_class_rep \<lbrakk>r\<rbrakk>"
and "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)"
and "tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<cong> r"
proof -
have 1: "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)"
proof -
have "\<exists>f g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) f g"
using assms BS1 isomorphic_symmetric rep_iso_class isomorphic_transitive
by blast
hence "isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) (tab\<^sub>0 r) (tab\<^sub>1 r)"
using assms tab\<^sub>0_def tab\<^sub>1_def
someI_ex [of "\<lambda>f. \<exists>g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) f g"]
someI_ex [of "\<lambda>g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) (tab\<^sub>0 r) g"]
by simp
thus ?thesis
using assms isomorphic_symmetric isomorphic_transitive rep_iso_class by blast
qed
hence "\<exists>\<phi>. \<guillemotleft>\<phi> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"
using isomorphic_def by blast
hence 2: "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso (rep r)"
using someI_ex [of "\<lambda>\<phi>. \<guillemotleft>\<phi> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"] rep_def by auto
show "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>"
using 2 by simp
show "iso (rep r)"
using 2 by simp
show "r \<cong> iso_class_rep \<lbrakk>r\<rbrakk>"
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 \<star> (tab\<^sub>0 r)\<^sup>* \<cong> r"
by simp
qed
lemma tab\<^sub>0_in_hom [intro]:
assumes "ide r"
shows "\<guillemotleft>tab\<^sub>0 r : src (tab\<^sub>0 r) \<rightarrow> src r\<guillemotright>"
and "\<guillemotleft>tab\<^sub>0 r : tab\<^sub>0 r \<Rightarrow> tab\<^sub>0 r\<guillemotright>"
proof -
show "\<guillemotleft>tab\<^sub>0 r : tab\<^sub>0 r \<Rightarrow> tab\<^sub>0 r\<guillemotright>"
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 "\<guillemotleft>tab\<^sub>0 r : src (tab\<^sub>0 r) \<rightarrow> src r\<guillemotright>"
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 "\<guillemotleft>tab\<^sub>1 r : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
and "\<guillemotleft>tab\<^sub>1 r : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 r\<guillemotright>"
proof -
show "\<guillemotleft>tab\<^sub>1 r : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 r\<guillemotright>"
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 "\<guillemotleft>tab\<^sub>1 r : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
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 "\<guillemotleft>rep r : src r \<rightarrow> trg r\<guillemotright>"
and "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>"
proof -
show "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>"
using assms rep_props by auto
thus "\<guillemotleft>rep r : src r \<rightarrow> trg r\<guillemotright>"
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 \<star> (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 \<open>
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 \<open>tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*\<close> its canonical tabulation, obtained as the adjoint
transpose of the identity, and then translate this to a tabulation of \<open>r\<close> via the chosen
isomorphism \<open>\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>\<close>.
\<close>
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 \<a> \<i> src trg \<open>tab\<^sub>0 r\<close>
using is_ide rep_props by unfold_locales auto
interpretation tab\<^sub>1: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>1 r\<close>
using is_ide rep_props by unfold_locales auto
text \<open>
A tabulation \<open>(tab\<^sub>0 r, tab, tab\<^sub>1 r)\<close> of \<open>r\<close> can be obtained as the adjoint transpose
of the isomorphism \<open>\<guillemotleft>rep r : (tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>\<close>. 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.
\<close>
definition tab
where "tab \<equiv> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) (rep r)"
text \<open>
In view of \<open>BS2'\<close>, the 1-cell \<open>(tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*\<close> has the canonical tabulation
obtained via adjoint transpose of an identity. In fact, this tabulation generates the
chosen tabulation of \<open>r\<close> in the same isomorphism class by translation along the
isomorphism \<open>\<guillemotleft>rep r : (tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>\<close>. This fact is used to show that the
mapping from 2-cells to arrows of spans preserves identities.
\<close>
lemma canonical_tabulation:
shows "tabulation V H \<a> \<i> src trg
((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) (tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)) (tab\<^sub>0 r) (tab\<^sub>1 r)"
proof -
have "\<exists>\<rho>. tabulation V H \<a> \<i> src trg ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) \<rho> (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 \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
and "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> tab = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have "tab = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) (rep r \<cdot> ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*))"
using tab_def is_ide rep_in_hom [of r] comp_arr_dom by auto
also have "... = (rep r \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
using is_ide tab\<^sub>0.trnr\<^sub>\<eta>_comp by auto
finally show 1: "tab = (rep r \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)" by simp
have "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> tab =
((inv (rep r) \<star> tab\<^sub>0 r) \<cdot> (rep r \<star> tab\<^sub>0 r)) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
unfolding 1 using comp_assoc by presburger
also have "... = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have 1: "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> (rep r \<star> tab\<^sub>0 r) = ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) \<star> 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 "\<guillemotleft>tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) :
tab\<^sub>1 r \<Rightarrow> (tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*) \<star> tab\<^sub>0 r\<guillemotright>"
by (meson canonical_tabulation tabulation_data.tab_in_hom(2) tabulation_def)
hence "((tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*) \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) =
tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (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) \<star> tab\<^sub>0 r) \<cdot> tab = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
by blast
qed
lemma tab_is_tabulation:
shows "tabulation V H \<a> \<i> 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 "\<guillemotleft>tab : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
and "\<guillemotleft>tab : tab\<^sub>1 r \<Rightarrow> r \<star> tab\<^sub>0 r\<guillemotright>"
proof -
interpret tab: tabulation V H \<a> \<i> src trg r tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
using tab_is_tabulation by simp
show "\<guillemotleft>tab : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
using tab.tab_in_hom by auto
show "\<guillemotleft>tab : tab\<^sub>1 r \<Rightarrow> r \<star> tab\<^sub>0 r\<guillemotright>"
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 \<star> tab\<^sub>0 r"
using tab_in_hom by auto
end
text \<open>
The following makes the chosen tabulation conveniently available whenever we are
considering a particular 1-cell.
\<close>
sublocale identity_in_bicategory_of_spans \<subseteq> tabulation V H \<a> \<i> src trg r tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
using is_ide tab_is_tabulation by simp
context identity_in_bicategory_of_spans
begin
interpretation tab\<^sub>0: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>0 r\<close>
using is_ide rep_props by unfold_locales auto
interpretation tab\<^sub>1: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>1 r\<close>
using is_ide rep_props by unfold_locales auto
text \<open>
The fact that adjoint transpose is a bijection allows us to invert the definition
of \<open>tab\<close> in terms of \<open>rep\<close> to express rep in terms of tab.
\<close>
lemma rep_in_terms_of_tab:
shows "rep r = T0.trnr\<^sub>\<epsilon> 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 \<open>
``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.
\<close>
lemma has_tabulation:
shows "\<exists>\<rho> f g. is_left_adjoint f \<and> is_left_adjoint g \<and> tabulation V H \<a> \<i> src trg r \<rho> 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 \<equiv> identity_in_bicategory_of_spans.tab V H \<a> \<i> src trg r"
abbreviation prj\<^sub>0
where "prj\<^sub>0 h k \<equiv> tab\<^sub>0 (k\<^sup>* \<star> h)"
abbreviation prj\<^sub>1
where "prj\<^sub>1 h k \<equiv> tab\<^sub>1 (k\<^sup>* \<star> h)"
lemma prj_in_hom [intro]:
assumes "ide h" and "is_left_adjoint k" and "trg h = trg k"
shows "\<guillemotleft>prj\<^sub>0 h k : src (prj\<^sub>0 h k) \<rightarrow> src h\<guillemotright>"
and "\<guillemotleft>prj\<^sub>1 h k : src (prj\<^sub>0 h k) \<rightarrow> src k\<guillemotright>"
and "\<guillemotleft>prj\<^sub>0 h k : prj\<^sub>0 h k \<Rightarrow> prj\<^sub>0 h k\<guillemotright>"
and "\<guillemotleft>prj\<^sub>1 h k : prj\<^sub>1 h k \<Rightarrow> prj\<^sub>1 h k\<guillemotright>"
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 \<open>
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.
\<close>
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 \<open>
The purpose of the somewhat strange-looking assumptions in this locale is
to cater to the form of data that we obtain from \<open>T1\<close>. Under the assumption
that we are in a bicategory of spans and that the legs of \<open>r\<close> and \<open>s\<close> are maps,
the hypothesized 2-cells will be uniquely determined isomorphisms, and an
arrow of spans \<open>w\<close> from \<open>r\<close> to \<open>s\<close> will be a map. We want to prove this once and
for all under the weakest assumptions we can manage.
\<close>
locale arrow_of_spans_of_maps =
bicategory_of_spans V H \<a> \<i> src trg +
r: span_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 +
s: span_of_maps V H \<a> \<i> src trg s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> '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: "\<exists>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
and leg1_iso: "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>"
begin
notation isomorphic (infix "\<cong>" 50)
lemma composite_leg1_is_map:
shows "is_left_adjoint (s\<^sub>1 \<star> 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 \<star> 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 \<star> w \<cong> r\<^sub>0"
and "\<exists>!\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
proof -
show 1: "s\<^sub>0 \<star> w \<cong> r\<^sub>0"
using leg0_lax composite_with_leg0_is_map r.leg0_is_map BS3 [of "s\<^sub>0 \<star> w" r\<^sub>0]
isomorphic_def
by auto
have "\<exists>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> iso \<theta>"
using 1 isomorphic_def by simp
moreover have "\<And>\<theta> \<theta>'. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<Longrightarrow> \<guillemotleft>\<theta>' : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<Longrightarrow> \<theta> = \<theta>'"
using BS3 r.leg0_is_map composite_with_leg0_is_map by blast
ultimately show "\<exists>!\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>" by blast
qed
lemma leg1_uniquely_isomorphic:
shows "r\<^sub>1 \<cong> s\<^sub>1 \<star> w"
and "\<exists>!\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
proof -
show 1: "r\<^sub>1 \<cong> s\<^sub>1 \<star> w"
using leg1_iso isomorphic_def by auto
have "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>"
using leg1_iso isomorphic_def isomorphic_symmetric by simp
moreover have "\<And>\<nu> \<nu>'. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<Longrightarrow> \<guillemotleft>\<nu>' : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<Longrightarrow> \<nu> = \<nu>'"
using BS3 r.leg1_is_map composite_leg1_is_map by blast
ultimately show "\<exists>!\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>" by blast
qed
definition the_\<theta>
where "the_\<theta> \<equiv> THE \<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
definition the_\<nu>
where "the_\<nu> \<equiv> THE \<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
lemma the_\<theta>_props:
shows "\<guillemotleft>the_\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>" and "iso the_\<theta>"
proof -
show "\<guillemotleft>the_\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
unfolding the_\<theta>_def
using the1I2 [of "\<lambda>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>" "\<lambda>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"]
leg0_uniquely_isomorphic
by simp
thus "iso the_\<theta>"
using BS3 r.leg0_is_map composite_with_leg0_is_map by simp
qed
lemma the_\<theta>_in_hom [intro]:
shows "\<guillemotleft>the_\<theta> : src r\<^sub>0 \<rightarrow> trg r\<^sub>0\<guillemotright>"
and "\<guillemotleft>the_\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
using the_\<theta>_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_\<theta>_simps [simp]:
shows "arr the_\<theta>"
and "src the_\<theta> = src r\<^sub>0" and "trg the_\<theta> = trg r\<^sub>0"
and "dom the_\<theta> = s\<^sub>0 \<star> w" and "cod the_\<theta> = r\<^sub>0"
using the_\<theta>_in_hom by auto
lemma the_\<nu>_props:
shows "\<guillemotleft>the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>" and "iso the_\<nu>"
proof -
show "\<guillemotleft>the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
unfolding the_\<nu>_def
using the1I2 [of "\<lambda>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>" "\<lambda>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"]
leg1_uniquely_isomorphic
by simp
thus "iso the_\<nu>"
using BS3 r.leg1_is_map composite_leg1_is_map by simp
qed
lemma the_\<nu>_in_hom [intro]:
shows "\<guillemotleft>the_\<nu> : src r\<^sub>1 \<rightarrow> trg r\<^sub>1\<guillemotright>"
and "\<guillemotleft>the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
using the_\<nu>_props apply auto
by (metis in_hhom_def in_homE isomorphic_implies_hpar(3) leg1_uniquely_isomorphic(1)
src_cod trg_dom)
lemma the_\<nu>_simps [simp]:
shows "arr the_\<nu>"
and "src the_\<nu> = src r\<^sub>1" and "trg the_\<nu> = trg r\<^sub>1"
and "dom the_\<nu> = r\<^sub>1" and "cod the_\<nu> = s\<^sub>1 \<star> w"
using the_\<nu>_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 \<a> \<i> src trg +
arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w +
\<sigma>: tabulation_data V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and w :: 'a
text \<open>
The following declaration allows us to inherit the rules and other facts defined in
locale @{locale uw\<theta>}. It is tedious to prove very much without these in place.
\<close>
sublocale arrow_of_spans_of_maps_to_tabulation_data \<subseteq> uw\<theta> V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 r\<^sub>0 w the_\<theta>
using \<sigma>.tab_in_hom is_ide the_\<theta>_props by unfold_locales auto
locale arrow_of_spans_of_maps_to_tabulation =
arrow_of_spans_of_maps_to_tabulation_data +
tabulation V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
locale tabulation_in_maps =
span_of_maps V H \<a> \<i> src trg s\<^sub>0 s\<^sub>1 +
tabulation V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
sublocale tabulation_in_maps \<subseteq> tabulation V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 ..
sublocale identity_in_bicategory_of_spans \<subseteq>
tabulation_in_maps V H \<a> \<i> src trg r tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
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 \<open>
The following sublocale declaration is perhaps pushing the limits of sensibility,
but the purpose is, given a cospan of maps \<open>(h, k)\<close>, to obtain ready access to the
composite \<open>k\<^sup>* \<star> h\<close> and its chosen tabulation.
\<close>
sublocale identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close>
using h_is_map k_is_map cospan left_adjoint_is_ide
by unfold_locales auto
notation isomorphic (infix "\<cong>" 50)
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
interpretation h: map_in_bicategory V H \<a> \<i> src trg h
using h_is_map by unfold_locales auto
interpretation k: map_in_bicategory V H \<a> \<i> src trg k
using k_is_map by unfold_locales auto
text \<open>
Our goal here is to reformulate the biuniversal properties of the chosen tabulation
of \<open>k\<^sup>* \<star> h\<close> in terms of its transpose, which yields a 2-cell from \<open>k \<star> tab\<^sub>1 (k\<^sup>* \<star> h)\<close>
to \<open>h \<star> tab\<^sub>0 (k\<^sup>* \<star> h)\<close>. These results do not depend on \<open>BS3\<close>.
\<close>
abbreviation p\<^sub>0
where "p\<^sub>0 \<equiv> prj\<^sub>0 h k"
abbreviation p\<^sub>1
where "p\<^sub>1 \<equiv> prj\<^sub>1 h k"
lemma p\<^sub>0_in_hom [intro]:
shows "\<guillemotleft>p\<^sub>0 : src p\<^sub>0 \<rightarrow> src h\<guillemotright>"
by auto
lemma p\<^sub>1_in_hom [intro]:
shows "\<guillemotleft>p\<^sub>1 : src p\<^sub>0 \<rightarrow> src k\<guillemotright>"
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 \<phi>
where "\<phi> \<equiv> k.trnl\<^sub>\<epsilon> (h \<star> p\<^sub>0) (\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> tab)"
lemma \<phi>_in_hom [intro]:
shows "\<guillemotleft>\<phi> : src p\<^sub>0 \<rightarrow> trg h\<guillemotright>"
and "\<guillemotleft>\<phi> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<phi> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright>"
unfolding \<phi>_def
using k.antipar cospan k.adjoint_transpose_left(2) [of "h \<star> p\<^sub>0" "p\<^sub>1"]
by fastforce
show "\<guillemotleft>\<phi> : src p\<^sub>0 \<rightarrow> trg h\<guillemotright>"
using 1 k.antipar arrI cospan vconn_implies_hpar(1-2) by force
qed
lemma \<phi>_simps [simp]:
shows "arr \<phi>"
and "src \<phi> = src p\<^sub>0" and "trg \<phi> = trg h"
and "dom \<phi> = k \<star> p\<^sub>1" and "cod \<phi> = h \<star> p\<^sub>0"
using \<phi>_in_hom by auto
lemma transpose_\<phi>:
shows "tab = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> k.trnl\<^sub>\<eta> p\<^sub>1 \<phi>"
proof -
have "\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> k.trnl\<^sub>\<eta> p\<^sub>1 \<phi> = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> tab"
unfolding \<phi>_def
using k.antipar cospan
k.adjoint_transpose_left(4)
[of "h \<star> p\<^sub>0" "p\<^sub>1" "\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> tab"]
by fastforce
also have "... = (\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0]) \<cdot> 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 "\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>" and "\<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright>"
shows "k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>) =
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>)"
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 \<and> 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 \<theta>"
using assms 1 k.antipar by auto
define \<chi> where "\<chi> = \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)"
have \<chi>: "\<guillemotleft>\<chi> : p\<^sub>1 \<star> w \<Rightarrow> k\<^sup>* \<star> h \<star> p\<^sub>0 \<star> w\<guillemotright>"
unfolding \<chi>_def
using assms 0 k.antipar cospan by (intro comp_in_homI, auto)
have "k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>) =
k.trnl\<^sub>\<epsilon> (h \<star> u) ((k\<^sup>* \<star> h \<star> \<theta>) \<cdot> \<chi> \<cdot> \<nu>)"
unfolding \<chi>_def
using assms 1 k.antipar cospan assoc_naturality [of "k\<^sup>*" h \<theta>] comp_assoc
by (metis "4" h.ide_left ide_char in_homE k.ide_right)
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) (k\<^sup>* \<star> h \<star> \<theta>) \<cdot> (k \<star> \<chi> \<cdot> \<nu>)"
proof -
have "ide (h \<star> u)"
using "1" u assms h.ide_left by blast
moreover have "seq (k\<^sup>* \<star> h \<star> \<theta>) (\<chi> \<cdot> \<nu>)"
using assms 1 k.antipar cospan \<chi> seqI'
apply (intro seqI)
apply auto
apply blast
proof -
have "dom (k\<^sup>* \<star> h \<star> \<theta>) = k\<^sup>* \<star> h \<star> p\<^sub>0 \<star> 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 \<chi>"
using \<chi> by auto
finally show "dom (k\<^sup>* \<star> h \<star> \<theta>) = cod \<chi>" by simp
qed
moreover have "src k = trg (k\<^sup>* \<star> h \<star> \<theta>)"
using assms k.antipar cospan calculation(2) by auto
ultimately show ?thesis
using k.trnl\<^sub>\<epsilon>_comp by simp
qed
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) (k\<^sup>* \<star> h \<star> \<theta>) \<cdot> (k \<star> \<chi>) \<cdot> (k \<star> \<nu>)"
using assms u \<chi> whisker_left
by (metis k.ide_left seqI')
also have
"... = (\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> u] \<cdot> (k \<star> k\<^sup>* \<star> h \<star> \<theta>)) \<cdot> (k \<star> \<chi>) \<cdot> (k \<star> \<nu>)"
unfolding k.trnl\<^sub>\<epsilon>_def by simp
also have "... = (h \<star> \<theta>) \<cdot>
(\<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot> (k \<star> \<chi>)) \<cdot>
(k \<star> \<nu>)"
proof -
have "\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> u] \<cdot> (k \<star> k\<^sup>* \<star> h \<star> \<theta>) =
\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
using assms 4 k.antipar cospan assoc'_naturality [of k "k\<^sup>*" "h \<star> \<theta>"]
by fastforce
also have "... = \<l>[h \<star> u] \<cdot> ((k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>)) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
using comp_assoc by presburger
also have "... = (\<l>[h \<star> u] \<cdot> (trg k \<star> h \<star> \<theta>)) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
proof -
have "(k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>) = k.\<epsilon> \<cdot> (k \<star> k\<^sup>*) \<star> (h \<star> u) \<cdot> (h \<star> \<theta>)"
using assms 1 k.antipar cospan interchange comp_arr_dom comp_cod_arr
by fastforce
also have "... = k.\<epsilon> \<star> h \<star> \<theta>"
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 \<star> h \<star> \<theta>) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w)"
using assms 4 k.antipar cospan whisker_left comp_arr_dom comp_cod_arr
interchange [of "trg k" k.\<epsilon> "h \<star> \<theta>" "h \<star> p\<^sub>0 \<star> w"]
by auto
finally have "(k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>) = (trg k \<star> h \<star> \<theta>) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (h \<star> \<theta>) \<cdot> \<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
proof -
have "\<l>[h \<star> u] \<cdot> (trg k \<star> h \<star> \<theta>) = (h \<star> \<theta>) \<cdot> \<l>[h \<star> p\<^sub>0 \<star> w]"
using assms 1 4 k.antipar cospan lunit_naturality [of "h \<star> \<theta>"]
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 "\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> u] \<cdot> (k \<star> k\<^sup>* \<star> h \<star> \<theta>) =
(h \<star> \<theta>) \<cdot> \<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>)"
proof -
have "\<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)) =
\<a>[h, p\<^sub>0, w] \<cdot> \<l>[(h \<star> p\<^sub>0) \<star> w] \<cdot>
(trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
proof -
have "\<l>[h \<star> p\<^sub>0 \<star> w] =
\<a>[h, p\<^sub>0, w] \<cdot> \<l>[(h \<star> p\<^sub>0) \<star> w] \<cdot> (trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
proof -
have "\<a>[h, p\<^sub>0, w] \<cdot> \<l>[(h \<star> p\<^sub>0) \<star> w] \<cdot> (trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) =
\<a>[h, p\<^sub>0, w] \<cdot> \<ll> ((h \<star> p\<^sub>0) \<star> w) \<cdot> (trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
using assms 0 k.antipar cospan comp_cod_arr \<ll>_ide_simp by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> \<ll> (\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
using assms 0 k.antipar cospan \<ll>.is_natural_2 [of "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w] \<cdot> \<ll> (h \<star> p\<^sub>0 \<star> w)"
using assms 0 k.antipar cospan \<ll>.is_natural_1 [of "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] comp_assoc
by simp
also have "... = (\<a>[h, p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<ll> (h \<star> p\<^sub>0 \<star> w)"
using comp_assoc by presburger
also have "... = \<ll> (h \<star> p\<^sub>0 \<star> w)"
using assms 0 k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp
also have "... = \<l>[h \<star> p\<^sub>0 \<star> w]"
using assms 0 k.antipar cospan \<ll>_ide_simp by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[trg h, h \<star> p\<^sub>0, w] \<cdot>
((trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w)) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
using assms 0 k.antipar cospan lunit_hcomp comp_assoc by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
(\<a>\<^sup>-\<^sup>1[trg h, h \<star> p\<^sub>0, w] \<cdot> (k.\<epsilon> \<star> (h \<star> p\<^sub>0) \<star> w)) \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
proof -
have "(trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) =
(k.\<epsilon> \<star> (h \<star> p\<^sub>0) \<star> w) \<cdot> ((k \<star> k\<^sup>*) \<star> \<a>\<^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.\<epsilon> "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]" "h \<star> p\<^sub>0 \<star> w"]
interchange [of k.\<epsilon> "k \<star> k\<^sup>*" "(h \<star> p\<^sub>0) \<star> w" "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> (h \<star> p\<^sub>0)) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
using assms 0 k.antipar cospan assoc'_naturality [of k.\<epsilon> "h \<star> p\<^sub>0" w] comp_assoc
by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> (h \<star> p\<^sub>0)) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot> (k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot>
(k \<star> tab \<star> w)"
proof -
have "k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) =
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot> (k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot>
(k \<star> tab \<star> w)"
proof -
have "seq \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] (\<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
using \<chi>_def assms 0 k.antipar cospan \<chi> 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 "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> (h \<star> p\<^sub>0)) \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot> (k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot>
\<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w]) \<cdot> ((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
proof -
have "k \<star> tab \<star> w = \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> (k \<star> tab \<star> w)"
proof -
have "\<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> (k \<star> tab \<star> w) =
(\<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w]) \<cdot> (k \<star> tab \<star> w)"
using comp_assoc by presburger
also have "... = (k \<star> ((k\<^sup>* \<star> h) \<star> p\<^sub>0) \<star> w) \<cdot> (k \<star> tab \<star> w)"
using assms k.antipar 0 comp_assoc_assoc' by simp
also have "... = k \<star> tab \<star> w"
using assms k.antipar 0 comp_cod_arr by simp
finally show ?thesis by simp
qed
also have "... = \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> ((k \<star> tab) \<star> w) \<cdot> \<a>\<^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 \<star> tab \<star> w = \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> ((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> h \<star> p\<^sub>0) \<star> w) \<cdot>
(\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<star> w) \<cdot>
((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot> ((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot>
(k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot> \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] =
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<star> w"
proof -
have "\<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot> ((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot>
(k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot> \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> ((\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle>, \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle>, (\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using assms 0 k.antipar cospan \<alpha>_def \<a>'_def by simp
also have "... = \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle>, \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<rbrace>"
using assms 0 k.antipar cospan
by (intro E.eval_eqI, simp_all)
also have "... = \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<star> w"
using assms 0 k.antipar cospan \<alpha>_def \<a>'_def by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot>
(\<l>[h \<star> p\<^sub>0] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<cdot>
(k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
using assms 0 k.antipar cospan comp_assoc whisker_right by auto
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
unfolding \<phi>_def k.trnl\<^sub>\<epsilon>_def
using assms 0 k.antipar cospan comp_assoc whisker_left by simp
finally have "\<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)) =
\<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
by blast
thus ?thesis
using \<chi>_def comp_assoc by simp
qed
finally show ?thesis by simp
qed
text \<open>
\<open>BS3\<close> implies that \<open>\<phi>\<close> is the unique 2-cell from \<open>k \<star> p\<^sub>1\<close> to \<open>h \<star> p\<^sub>0\<close> and is an isomorphism.
\<close>
lemma \<phi>_uniqueness:
shows "\<And>\<mu>. \<guillemotleft>\<mu> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright> \<Longrightarrow> \<mu> = \<phi>" and "iso \<phi>"
proof -
have 2: "is_left_adjoint (k \<star> p\<^sub>1)"
using k.antipar cospan left_adjoints_compose by (simp add: k_is_map)
have 3: "is_left_adjoint (h \<star> p\<^sub>0)"
using k.antipar cospan left_adjoints_compose by (simp add: h_is_map)
show "\<And>\<mu>. \<guillemotleft>\<mu> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright> \<Longrightarrow> \<mu> = \<phi>"
using \<phi>_in_hom 2 3 BS3 by simp
show "iso \<phi>"
using \<phi>_in_hom 2 3 BS3 by simp
qed
text \<open>
As a consequence, the chosen tabulation of \<open>k\<^sup>* \<star> h\<close> is the unique 2-cell from
\<open>p\<^sub>1\<close> to \<open>(k\<^sup>* \<star> h) \<star> p\<^sub>0\<close>, and therefore if we are given any such 2-cell we may
conclude it yields a tabulation of \<open>k\<^sup>* \<star> h\<close>.
\<close>
lemma tab_uniqueness:
assumes "\<guillemotleft>\<tau> : p\<^sub>1 \<Rightarrow> (k\<^sup>* \<star> h) \<star> p\<^sub>0\<guillemotright>"
shows "\<tau> = tab"
proof -
have "\<guillemotleft>k.trnl\<^sub>\<epsilon> (h \<star> p\<^sub>0) (\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> \<tau>) : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright>"
using assms k.antipar cospan k.adjoint_transpose_left(2) [of "h \<star> p\<^sub>0" "p\<^sub>1"]
assoc_in_hom
by force
hence "tab = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> k.trnl\<^sub>\<eta> p\<^sub>1 (k.trnl\<^sub>\<epsilon> (h \<star> p\<^sub>0) (\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> \<tau>))"
using transpose_\<phi> \<phi>_uniqueness(1) by auto
also have "... = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> \<tau>"
using assms k.antipar cospan k.adjoint_transpose_left(4) assoc_in_hom by auto
also have "... = (\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0]) \<cdot> \<tau>"
using comp_assoc by presburger
also have "... = \<tau>"
using assms k.antipar cospan comp_cod_arr comp_assoc_assoc' by auto
finally show ?thesis by simp
qed
text \<open>
The following lemma reformulates the biuniversal property of the canonical tabulation
of \<open>k\<^sup>* \<star> h\<close> as a biuniversal property of \<open>\<phi>\<close>, regarded as a square that commutes up to
isomorphism.
\<close>
lemma \<phi>_biuniversal_prop:
assumes "ide u" and "ide v"
shows "\<And>\<mu>. \<guillemotleft>\<mu> : k \<star> v \<Rightarrow> h \<star> u\<guillemotright> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
and "\<And>w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w';
\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>;
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] =
(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> \<beta>) \<rbrakk>
\<Longrightarrow> \<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<theta> = \<theta>' \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<beta>"
proof -
fix \<mu>
assume \<mu>: "\<guillemotleft>\<mu> : k \<star> v \<Rightarrow> h \<star> u\<guillemotright>"
have 1: "src h = trg u"
using assms \<mu> ide_cod
by (metis ide_def in_homE seq_if_composable)
have 2: "src k = trg v"
using assms \<mu> ide_dom
by (metis ideD(1) in_homE not_arr_null seq_if_composable)
let ?\<omega> = "\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, u] \<cdot> k.trnl\<^sub>\<eta> v \<mu>"
have \<omega>: "\<guillemotleft>?\<omega> : v \<Rightarrow> (k\<^sup>* \<star> h) \<star> u\<guillemotright>"
using assms \<mu> 1 2 k.antipar cospan k.adjoint_transpose_left(1) [of "h \<star> u" v]
assoc_in_hom
by auto
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu> = ?\<omega>"
using assms \<omega> T1 [of u ?\<omega>] comp_assoc by (metis in_homE)
have 0: "src p\<^sub>0 = trg w"
using w\<theta>\<nu> ide_dom
by (metis hseqE ideD(1) in_homE)
have "\<mu> = k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>)"
proof -
have "\<mu> = k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ?\<omega>)"
proof -
have "k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ?\<omega>) =
k.trnl\<^sub>\<epsilon> (h \<star> u) ((\<a>[k\<^sup>*, h, u] \<cdot> \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \<cdot> k.trnl\<^sub>\<eta> v \<mu>)"
using comp_assoc by presburger
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) (k.trnl\<^sub>\<eta> v \<mu>)"
proof -
have "(\<a>[k\<^sup>*, h, u] \<cdot> \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \<cdot> k.trnl\<^sub>\<eta> v \<mu> = (k\<^sup>* \<star> h \<star> u) \<cdot> k.trnl\<^sub>\<eta> v \<mu>"
using comp_assoc_assoc'
by (simp add: "1" assms(1) cospan k.antipar(2))
also have "... = k.trnl\<^sub>\<eta> v \<mu>"
using "1" \<omega> assms(1) comp_ide_arr cospan k.antipar(2) by fastforce
finally show ?thesis
by simp
qed
also have "... = \<mu>"
using assms \<mu> k.antipar cospan 1 2 k.adjoint_transpose_left(3) by simp
finally show ?thesis by simp
qed
thus ?thesis using w\<theta>\<nu> by simp
qed
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>)"
using assms k.antipar cospan w\<theta>\<nu> transpose_triangle [of w \<theta> u \<nu>] by auto
finally have "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
by simp
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
using w\<theta>\<nu> by blast
next
fix w w' \<theta> \<theta>' \<beta>
assume w: "ide w"
assume w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>"
assume eq: "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] =
(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> \<beta>)"
have 0: "src p\<^sub>0 = trg w"
using \<theta> ide_dom
by (metis ideD(1) in_homE not_arr_null seq_if_composable)
interpret uw\<theta>w'\<theta>': uw\<theta>w'\<theta>' V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close>
u w \<theta> w' \<theta>'
using w \<theta> w' \<theta>' apply (unfold_locales) by auto
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<theta> = \<theta>' \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<beta>"
proof -
let ?LHS = "\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)"
let ?RHS = "\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>') \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w'] \<cdot> (tab \<star> w') \<cdot> \<beta>"
have eq': "?LHS = ?RHS"
proof -
have "k.trnl\<^sub>\<epsilon> (h \<star> u) ?LHS =
k.trnl\<^sub>\<epsilon> (h \<star> u)
(\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> (p\<^sub>1 \<star> w))"
using assms 0 w \<theta> \<beta> k.antipar cospan comp_arr_dom
by (metis tab_simps(1) tab_simps(4) whisker_right)
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> p\<^sub>1 \<star> w)"
using assms w \<theta> \<beta> 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 \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
using assms 0 w k.antipar cospan comp_arr_dom by simp
also have "... = (h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> \<beta>)"
using eq by blast
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) ?RHS"
using assms w' \<theta>' \<beta> transpose_triangle by simp
finally have 4: "k.trnl\<^sub>\<epsilon> (h \<star> u) ?LHS = k.trnl\<^sub>\<epsilon> (h \<star> u) ?RHS"
by simp
have "src k = trg (p\<^sub>1 \<star> w)"
using assms 0 w k.antipar cospan by simp
moreover have "src k\<^sup>* = trg (h \<star> u)"
using assms 0 w k.antipar cospan by simp
moreover have "ide (h \<star> u)"
using assms 0 w k.antipar cospan by simp
moreover have "ide (p\<^sub>1 \<star> w)"
using assms 0 w k.antipar cospan by simp
ultimately have "inj_on (k.trnl\<^sub>\<epsilon> (h \<star> u)) (hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u))"
using assms 0 w w' k.antipar cospan k.adjoint_transpose_left(6) bij_betw_imp_inj_on
by blast
moreover have "?LHS \<in> hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u)"
proof -
have "\<guillemotleft>\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) :
p\<^sub>1 \<star> w \<Rightarrow> k\<^sup>* \<star> h \<star> u\<guillemotright>"
using k.antipar cospan
apply (intro comp_in_homI) by auto
thus ?thesis by simp
qed
moreover have "?RHS \<in> hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u)"
proof -
have "\<guillemotleft>\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>') \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w'] \<cdot>
(tab \<star> w') \<cdot> \<beta> : p\<^sub>1 \<star> w \<Rightarrow> k\<^sup>* \<star> h \<star> u\<guillemotright>"
using \<beta> 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>\<epsilon> (h \<star> u)" "hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u)"]
by simp
qed
moreover have "seq \<a>[k\<^sup>*, h, u] (composite_cell w \<theta>)"
using assms k.antipar cospan tab_in_hom by fastforce
moreover have "seq \<a>[k\<^sup>*, h, u] (composite_cell w' \<theta>' \<cdot> \<beta>)"
using assms \<beta> k.antipar cospan tab_in_hom by fastforce
ultimately have "composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta>"
using assms 0 w w' \<beta> k.antipar cospan iso_assoc iso_is_section section_is_mono
monoE [of "\<a>[k\<^sup>*, h, u]" "composite_cell w \<theta>" "composite_cell w' \<theta>' \<cdot> \<beta>"]
comp_assoc
by simp
thus ?thesis
using w w' \<theta> \<theta>' \<beta> eq' T2 [of w w' \<theta> u \<theta>' \<beta>] by metis
qed
qed
text \<open>
Using the uniqueness properties established for \<open>\<phi>\<close>, we obtain yet another reformulation
of the biuniversal property associated with the chosen tabulation of \<open>k\<^sup>* \<star> h\<close>,
this time as a kind of pseudo-pullback. We will use this to show that the
category of isomorphism classes of maps has pullbacks.
\<close>
lemma has_pseudo_pullback:
assumes "is_left_adjoint u" and "is_left_adjoint v" and "isomorphic (k \<star> v) (h \<star> u)"
shows "\<exists>w. is_left_adjoint w \<and> isomorphic (p\<^sub>0 \<star> w) u \<and> isomorphic v (p\<^sub>1 \<star> w)"
and "\<And>w w'. \<lbrakk> is_left_adjoint w; is_left_adjoint w';
p\<^sub>0 \<star> w \<cong> u; v \<cong> p\<^sub>1 \<star> w; p\<^sub>0 \<star> w' \<cong> u; v \<cong> p\<^sub>1 \<star> w' \<rbrakk> \<Longrightarrow> w \<cong> w'"
proof -
interpret u: map_in_bicategory V H \<a> \<i> src trg u
using assms(1) by unfold_locales auto
interpret v: map_in_bicategory V H \<a> \<i> src trg v
using assms(2) by unfold_locales auto
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : k \<star> v \<Rightarrow> h \<star> u\<guillemotright> \<and> iso \<mu>"
using assms(3) by auto
obtain w \<theta> \<nu> where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
using assms \<mu> \<phi>_biuniversal_prop(1) [of u v \<mu>] by auto
have "is_left_adjoint w \<and> isomorphic (p\<^sub>0 \<star> w) u \<and> isomorphic v (p\<^sub>1 \<star> w)"
proof (intro conjI)
show 1: "is_left_adjoint w"
using assms(2) w\<theta>\<nu> left_adjoint_preserved_by_iso' isomorphic_def BS4 leg1_is_map
by blast
show "v \<cong> p\<^sub>1 \<star> w"
using w\<theta>\<nu> isomorphic_def by blast
show "p\<^sub>0 \<star> w \<cong> u"
proof -
have "src p\<^sub>0 = trg w"
using w\<theta>\<nu> ide_dom
by (metis ideD(1) in_homE not_arr_null seq_if_composable)
hence "is_left_adjoint (p\<^sub>0 \<star> w)"
using 1 left_adjoints_compose by simp
thus ?thesis
using assms w\<theta>\<nu> 1 BS3 isomorphic_def by metis
qed
qed
thus "\<exists>w. is_left_adjoint w \<and> p\<^sub>0 \<star> w \<cong> u \<and> v \<cong> p\<^sub>1 \<star> w"
by blast
show "\<And>w w'. \<lbrakk> is_left_adjoint w; is_left_adjoint w';
p\<^sub>0 \<star> w \<cong> u; v \<cong> p\<^sub>1 \<star> w; p\<^sub>0 \<star> w' \<cong> u; v \<cong> p\<^sub>1 \<star> w' \<rbrakk> \<Longrightarrow> w \<cong> w'"
proof -
fix w w'
assume w: "is_left_adjoint w" and w': "is_left_adjoint w'"
assume 1: "p\<^sub>0 \<star> w \<cong> u"
assume 2: "v \<cong> p\<^sub>1 \<star> w"
assume 3: "p\<^sub>0 \<star> w' \<cong> u"
assume 4: "v \<cong> p\<^sub>1 \<star> w'"
obtain \<theta> where \<theta>: "\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>"
using 1 by auto
obtain \<theta>' where \<theta>': "\<guillemotleft>\<theta>' : p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>"
using 3 by auto
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu>: v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>"
using 2 by blast
obtain \<nu>' where \<nu>': "\<guillemotleft>\<nu>': v \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>'"
using 4 by blast
let ?\<beta> = "\<nu>' \<cdot> inv \<nu>"
have \<beta>: "\<guillemotleft>?\<beta> : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>"
using \<nu> \<nu>' by (elim conjE in_homE, auto)
interpret uw\<theta>: uw\<theta> V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close> u w \<theta>
using w \<theta> left_adjoint_is_ide
by unfold_locales auto
interpret uw'\<theta>': uw\<theta> V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close>
u w' \<theta>'
using w' \<theta>' left_adjoint_is_ide
by unfold_locales auto
interpret uw\<theta>w'\<theta>': uw\<theta>w'\<theta>' V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close> u w \<theta> w' \<theta>'
using w w' \<theta> \<theta>' left_adjoint_is_ide by unfold_locales
have "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] =
(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot>
(k \<star> ?\<beta>)"
proof -
let ?LHS = "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
let ?RHS = "(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> ?\<beta>)"
have "\<guillemotleft>?LHS : k \<star> p\<^sub>1 \<star> w \<Rightarrow> h \<star> u\<guillemotright>"
using w k.antipar by fastforce
moreover have "\<guillemotleft>?RHS : k \<star> p\<^sub>1 \<star> w \<Rightarrow> h \<star> u\<guillemotright>"
using w k.antipar \<beta> by fastforce
moreover have "is_left_adjoint (k \<star> p\<^sub>1 \<star> w)"
using w k.is_map left_adjoints_compose by simp
moreover have "is_left_adjoint (h \<star> u)"
using assms h.is_map left_adjoints_compose by auto
ultimately show "?LHS = ?RHS"
using BS3 by blast
qed
hence "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<theta> = \<theta>' \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = ?\<beta>"
using assms left_adjoint_is_ide w w' \<theta> \<theta>' \<nu> \<nu>' \<beta>
\<phi>_biuniversal_prop(2) [of u v w w' \<theta> \<theta>' ?\<beta>]
by presburger
thus "w \<cong> w'"
using w w' BS3 isomorphic_def by metis
qed
qed
end
subsubsection "Tabulations in Maps"
text \<open>
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.
\<close>
context tabulation_in_maps
begin
text \<open>
The following are the conditions under which \<open>w\<close> is a 1-cell induced via \<open>T1\<close> by
a 2-cell \<open>\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> s \<star> r\<^sub>0\<guillemotright>\<close>: \<open>w\<close> is an arrow of spans and \<open>\<omega>\<close> is obtained by
composing the tabulation \<open>\<sigma>\<close> with \<open>w\<close> and the isomorphisms that witness \<open>w\<close> being
an arrow of spans.
\<close>
abbreviation is_induced_by_cell
where "is_induced_by_cell w r\<^sub>0 \<omega> \<equiv>
arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<omega>) s\<^sub>0 s\<^sub>1 w \<and>
composite_cell w (arrow_of_spans_of_maps.the_\<theta> V H r\<^sub>0 s\<^sub>0 w) \<cdot>
(arrow_of_spans_of_maps.the_\<nu> V H (dom \<omega>) s\<^sub>1 w) = \<omega>"
lemma induced_map_unique:
assumes "is_induced_by_cell w r\<^sub>0 \<omega>" and "is_induced_by_cell w' r\<^sub>0 \<omega>"
shows "isomorphic w w'"
proof -
interpret w: arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s\<^sub>0 s\<^sub>1 w
using assms(1) by auto
interpret w: arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w
..
interpret w': arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s\<^sub>0 s\<^sub>1 w'
using assms(2) by auto
interpret w': arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w'
..
let ?\<beta> = "w'.the_\<nu> \<cdot> inv w.the_\<nu>"
have \<beta>: "\<guillemotleft>?\<beta> : s\<^sub>1 \<star> w \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright>"
using w.the_\<nu>_props w'.the_\<nu>_props arr_iff_in_hom by auto
have 1: "composite_cell w w.the_\<theta> = composite_cell w' w'.the_\<theta> \<cdot> (w'.the_\<nu> \<cdot> inv w.the_\<nu>)"
proof -
have "composite_cell w' w'.the_\<theta> \<cdot> (w'.the_\<nu> \<cdot> inv w.the_\<nu>) =
((composite_cell w' w'.the_\<theta>) \<cdot> w'.the_\<nu>) \<cdot> inv w.the_\<nu>"
using comp_assoc by presburger
also have "... = \<omega> \<cdot> inv w.the_\<nu>"
using assms(2) comp_assoc by simp
also have "... = (composite_cell w w.the_\<theta> \<cdot> w.the_\<nu>) \<cdot> inv w.the_\<nu>"
using assms(1) comp_assoc by simp
also have "... = composite_cell w w.the_\<theta> \<cdot> w.the_\<nu> \<cdot> inv w.the_\<nu>"
using comp_assoc by presburger
also have "... = composite_cell w w.the_\<theta>"
proof -
have "w.the_\<nu> \<cdot> inv w.the_\<nu> = s\<^sub>1 \<star> w"
using w.the_\<nu>_props comp_arr_inv inv_is_inverse by auto
thus ?thesis
using composite_cell_in_hom w.ide_w w.the_\<theta>_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 "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
using 1 \<beta> w.is_ide w'.is_ide w.the_\<theta>_props w'.the_\<theta>_props
T2 [of w w' w.the_\<theta> r\<^sub>0 w'.the_\<theta> ?\<beta>]
by blast
thus ?thesis
using BS3 w.is_map w'.is_map by blast
qed
text \<open>
The object src \<open>s\<^sub>0\<close> forming the apex of the tabulation satisfies the conditions for
being a map induced via \<open>T1\<close> by the 2-cell \<open>\<sigma>\<close> itself. This is ultimately required
for the map from 2-cells to arrows of spans to be functorial with respect to vertical
composition.
\<close>
lemma apex_is_induced_by_cell:
shows "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \<sigma>"
proof -
have 1: "arrow_of_spans_of_maps V H \<a> \<i> 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 \<a> \<i> src trg s\<^sub>0 \<open>dom \<sigma>\<close> s\<^sub>0 s\<^sub>1 \<open>src s\<^sub>0\<close>
using 1 tab_in_hom by simp
interpret w: arrow_of_spans_of_maps_to_tabulation
V H \<a> \<i> src trg s\<^sub>0 \<open>dom \<sigma>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 \<open>src s\<^sub>0\<close>
..
show "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \<sigma>"
proof (intro conjI)
show "arrow_of_spans_of_maps V H \<a> \<i> src trg s\<^sub>0 (dom \<sigma>) 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_\<theta> \<cdot> w.the_\<nu> = \<sigma>"
proof -
have \<theta>: "w.the_\<theta> = \<r>[s\<^sub>0]"
using iso_runit [of s\<^sub>0] w.leg0_uniquely_isomorphic w.the_\<theta>_props
the1_equality [of "\<lambda>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> src s\<^sub>0 \<Rightarrow> s\<^sub>0\<guillemotright> \<and> iso \<theta>"]
by auto
have \<nu>: "w.the_\<nu> = \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using iso_runit' w.leg1_uniquely_isomorphic w.the_\<nu>_props leg1_simps(3)
the1_equality [of "\<lambda>\<nu>. \<guillemotleft>\<nu> : s\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> src s\<^sub>0\<guillemotright> \<and> iso \<nu>"] tab_in_vhom'
by auto
have "composite_cell (src s\<^sub>0) \<r>[s\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1] = \<sigma>"
proof -
have "composite_cell (src s\<^sub>0) \<r>[s\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1] =
((s \<star> \<r>[s\<^sub>0]) \<cdot> \<a>[s, s\<^sub>0, src s\<^sub>0]) \<cdot> (\<sigma> \<star> src s\<^sub>0) \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using comp_assoc by presburger
also have "... = (\<r>[s \<star> s\<^sub>0] \<cdot> (\<sigma> \<star> src s\<^sub>0)) \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using runit_hcomp comp_assoc by simp
also have "... = \<sigma> \<cdot> \<r>[s\<^sub>1] \<cdot> \<r>\<^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 "... = \<sigma>"
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 \<theta> \<nu> comp_assoc by simp
qed
qed
qed
end
subsubsection "Composing Tabulations"
text \<open>
Given tabulations \<open>(r\<^sub>0, \<rho>, r\<^sub>1)\<close> of \<open>r\<close> and \<open>(s\<^sub>0, \<sigma>, s\<^sub>1)\<close> of \<open>s\<close> in a bicategory of spans,
where \<open>(r\<^sub>0, r\<^sub>1)\<close> and \<open>(s\<^sub>0, s\<^sub>1)\<close> are spans of maps and 1-cells \<open>r\<close> and \<open>s\<close> are composable,
we can construct a 2-cell that yields a tabulation of \<open>r \<star> s\<close>.
The proof uses the fact that the 2-cell \<open>\<phi>\<close> associated with the cospan \<open>(r\<^sub>0, s\<^sub>1)\<close>
is an isomorphism, which we have proved above
(\<open>cospan_of_maps_in_bicategory_of_spans.\<phi>_uniqueness\<close>) using \<open>BS3\<close>.
However, this is the only use of \<open>BS3\<close> in the proof, and it seems plausible that it would be
possible to establish that \<open>\<phi>\<close> 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 \<open>\<phi>\<close> is an isomorphism as part of
a weakening of \<open>BS3\<close>.
\<close>
locale composite_tabulation_in_maps =
bicategory_of_spans V H \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a +
assumes composable: "src r = trg s"
begin
text \<open>
Interpret \<open>(r\<^sub>0, s\<^sub>1)\<close> as a @{locale cospan_of_maps_in_bicategory_of_spans},
to obtain the isomorphism \<open>\<phi>\<close> in the central diamond, along with the assertion
that it is unique.
\<close>
interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg s\<^sub>1 r\<^sub>0
using \<rho>.leg0_is_map \<sigma>.leg1_is_map composable by unfold_locales auto
text \<open>
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.
\<close>
lemma r\<^sub>0s\<^sub>1_is_cospan:
shows "cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg s\<^sub>1 r\<^sub>0"
..
text \<open>
The following define the projections associated with the natural tabulation of \<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close>.
\<close>
abbreviation p\<^sub>0
where "p\<^sub>0 \<equiv> r\<^sub>0s\<^sub>1.p\<^sub>0"
abbreviation p\<^sub>1
where "p\<^sub>1 \<equiv> r\<^sub>0s\<^sub>1.p\<^sub>1"
text \<open>
$$
\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}
}
$$
\<close>
text \<open>
Next, we define the 2-cell that is the composite of the tabulation \<open>\<sigma>\<close>, the tabulation \<open>\<rho>\<close>,
and the central diamond that commutes up to unique isomorphism \<open>\<phi>\<close>.
\<close>
definition tab
where "tab \<equiv> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1)"
lemma tab_in_hom [intro]:
shows "\<guillemotleft>tab : r\<^sub>1 \<star> p\<^sub>1 \<Rightarrow> (r \<star> s) \<star> s\<^sub>0 \<star> p\<^sub>0\<guillemotright>"
using \<rho>.T0.antipar(1) r\<^sub>0s\<^sub>1.\<phi>_in_hom composable \<rho>.leg0_in_hom(1) \<sigma>.leg1_in_hom(1)
composable tab_def
by auto
interpretation tabulation_data V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close>
using composable tab_in_hom
by unfold_locales auto
text \<open>
In the subsequent proof we will use coherence to shortcut a few of the calculations.
\<close>
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
text \<open>
The following is applied twice in the proof of property \<open>T2\<close> for the composite
tabulation. It's too long to repeat.
\<close>
lemma technical:
assumes "ide w"
and "\<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>"
and "w\<^sub>r = p\<^sub>1 \<star> w"
and "\<theta>\<^sub>r = (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
shows "\<rho>.composite_cell w\<^sub>r \<theta>\<^sub>r = \<a>[r, s, u] \<cdot> composite_cell w \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
text \<open>
$$
\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}
}
$$
\<close>
proof -
interpret uw\<theta>: uw\<theta> V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close> u w \<theta>
using assms(1-2) composable
by unfold_locales auto
show ?thesis
proof -
have "\<a>[r, s, u] \<cdot> composite_cell w \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] =
(\<a>[r, s, u] \<cdot> ((r \<star> s) \<star> \<theta>)) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = (r \<star> s \<star> \<theta>) \<cdot> \<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(tab \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using assoc_naturality [of r s \<theta>] composable comp_assoc by simp
also have "... = (r \<star> s \<star> \<theta>) \<cdot> \<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
((\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) \<cdot>
\<rho>.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
unfolding tab_def
using comp_assoc by presburger
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<cdot> \<rho>.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using composable \<rho>.T0.antipar(1) comp_assoc whisker_right by auto
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot>
(\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using composable \<rho>.T0.antipar(1) whisker_right tab_def tab_in_hom(2)
composable comp_assoc
by force
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot>
((\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]) \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
using assoc'_naturality [of \<rho> p\<^sub>1 w] \<rho>.T0.antipar(1) r\<^sub>0s\<^sub>1.base_simps(2) comp_assoc
by auto
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w]) \<cdot> \<rho>.composite_cell (p\<^sub>1 \<star> w) \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] =
\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"
proof -
have "(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) =
\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \<star> w]"
using pentagon' \<rho>.T0.antipar(1) comp_assoc by simp
moreover have 1: "seq (\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w)(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]))"
using \<rho>.T0.antipar(1)
by (intro seqI hseqI, auto)
ultimately
have "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] =
((\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"
using \<rho>.T0.antipar(1) iso_assoc
invert_side_of_triangle(2)
[of "(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])"
"\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]" "\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \<star> w]"]
by fastforce
hence "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] =
(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"
using comp_assoc by presburger
moreover have "seq (inv (\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w)) \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) 1 by fastforce
ultimately show ?thesis
using \<rho>.T0.antipar(1) iso_assoc
invert_side_of_triangle(1)
[of "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]" "\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w"
"\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"]
by fastforce
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
(((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, s\<^sub>1 \<star> p\<^sub>0, w]) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<rho>.composite_cell (p\<^sub>1 \<star> w) \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] = \<a>\<^sup>-\<^sup>1[r, s\<^sub>1 \<star> p\<^sub>0, w] \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using assoc'_naturality [of r r\<^sub>0s\<^sub>1.\<phi> w] r\<^sub>0s\<^sub>1.cospan by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot>
(\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w]) \<cdot>
(r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<rho>.composite_cell (p\<^sub>1 \<star> w) \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, s\<^sub>1 \<star> p\<^sub>0, w] =
\<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w)"
using assoc'_naturality [of r "\<sigma> \<star> p\<^sub>0" w]
by (simp add: composable)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot>
(r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
((r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
proof -
have "\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] =
r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]"
proof -
have "\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, (\<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, (\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<alpha>_def \<a>'_def composable by simp
also have "... = \<lbrace>\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using composable
by (intro E.eval_eqI, simp_all)
also have "... = r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]"
using \<alpha>_def \<a>'_def composable by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot>
(r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
\<rho>.composite_cell (p\<^sub>1 \<star> w)
(((\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])"
using assms(3) arrI \<rho>.T0.antipar(1) whisker_left by auto
also have "... = (r \<star> (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<a>[s \<star> s\<^sub>0, p\<^sub>0, w] \<cdot> ((\<sigma> \<star> p\<^sub>0) \<star> w)) \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
using \<rho>.T0.antipar(1) comp_assoc whisker_left by auto
also have "... = (r \<star> (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
using assoc_naturality [of \<sigma> 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 \<a> \<i> src trg (r \<star> s) tab (s\<^sub>0 \<star> p\<^sub>0) (r\<^sub>1 \<star> p\<^sub>1)"
proof
show "\<And>u \<omega>. \<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> (r \<star> s) \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
proof -
fix u \<omega>
assume u: "ide u"
assume \<omega>: "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> (r \<star> s) \<star> u\<guillemotright>"
let ?v = "dom \<omega>"
have 1: "\<guillemotleft>\<a>[r, s, u] \<cdot> \<omega> : ?v \<Rightarrow> r \<star> s \<star> u\<guillemotright>"
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 \<omega>.
* 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 \<and> src s = trg u"
by (metis \<omega> arr_cod hseq_char in_homE hcomp_simps(1))
thus ?thesis
using u \<omega> by fastforce
qed
obtain w\<^sub>r \<theta>\<^sub>r \<nu>\<^sub>r
where w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r: "ide w\<^sub>r \<and> \<guillemotleft>\<theta>\<^sub>r : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s \<star> u\<guillemotright> \<and>
\<guillemotleft>\<nu>\<^sub>r : ?v \<Rightarrow> r\<^sub>1 \<star> w\<^sub>r\<guillemotright> \<and> iso \<nu>\<^sub>r \<and>
\<rho>.composite_cell w\<^sub>r \<theta>\<^sub>r \<cdot> \<nu>\<^sub>r = \<a>[r, s, u] \<cdot> \<omega>"
using u \<omega> \<rho>.T1 [of "s \<star> u" "\<a>[r, s, u] \<cdot> \<omega>"]
by (metis 1 \<rho>.ide_base \<sigma>.ide_base arr_cod composable ide_hcomp in_homE
match_1 not_arr_null seq_if_composable)
text \<open>
$$
\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}
}
$$
\<close>
text \<open>We need some simps, etc., otherwise the subsequent diagram chase is very painful.\<close>
have w\<^sub>r: "ide w\<^sub>r"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
have [simp]: "src w\<^sub>r = src u"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<omega> 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 \<rho>"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r
by (metis 1 arrI not_arr_null seqE seq_if_composable)
have \<theta>\<^sub>r_in_hom [intro]: "\<guillemotleft>\<theta>\<^sub>r : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s \<star> u\<guillemotright>"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
have \<theta>\<^sub>r_in_hhom [intro]: "\<guillemotleft>\<theta>\<^sub>r : src u \<rightarrow> trg s\<guillemotright>"
using \<theta>\<^sub>r_in_hom src_cod [of \<theta>\<^sub>r] trg_cod [of \<theta>\<^sub>r]
by (metis \<open>src w\<^sub>r = src u\<close> \<sigma>.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 \<theta>\<^sub>r = src u" using \<theta>\<^sub>r_in_hhom by auto
have [simp]: "trg \<theta>\<^sub>r = trg s" using \<theta>\<^sub>r_in_hhom by auto
have [simp]: "dom \<theta>\<^sub>r = r\<^sub>0 \<star> w\<^sub>r" using \<theta>\<^sub>r_in_hom by blast
have [simp]: "cod \<theta>\<^sub>r = s \<star> u" using \<theta>\<^sub>r_in_hom by blast
have \<nu>\<^sub>r_in_hom [intro]: "\<guillemotleft>\<nu>\<^sub>r : ?v \<Rightarrow> r\<^sub>1 \<star> w\<^sub>r\<guillemotright>" using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
have \<nu>\<^sub>r_in_hhom [intro]: "\<guillemotleft>\<nu>\<^sub>r : src u \<rightarrow> trg r\<guillemotright>"
using \<nu>\<^sub>r_in_hom src_dom [of \<nu>\<^sub>r] trg_dom [of \<nu>\<^sub>r]
by (metis \<open>src w\<^sub>r = src u\<close> \<rho>.leg1_simps(4) arr_cod in_hhomI in_homE
src_hcomp trg_hcomp vconn_implies_hpar(3) vconn_implies_hpar(4))
have [simp]: "src \<nu>\<^sub>r = src u" using \<nu>\<^sub>r_in_hhom by auto
have [simp]: "trg \<nu>\<^sub>r = trg r" using \<nu>\<^sub>r_in_hhom by auto
have [simp]: "dom \<nu>\<^sub>r = ?v" using \<nu>\<^sub>r_in_hom by auto
have [simp]: "cod \<nu>\<^sub>r = r\<^sub>1 \<star> w\<^sub>r" using \<nu>\<^sub>r_in_hom by auto
have iso_\<nu>\<^sub>r: "iso \<nu>\<^sub>r" using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
obtain w\<^sub>s \<theta>\<^sub>s \<nu>\<^sub>s
where w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s: "ide w\<^sub>s \<and> \<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> w\<^sub>s \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu>\<^sub>s : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s\<^sub>1 \<star> w\<^sub>s\<guillemotright> \<and> iso \<nu>\<^sub>s \<and>
\<sigma>.composite_cell w\<^sub>s \<theta>\<^sub>s \<cdot> \<nu>\<^sub>s = \<theta>\<^sub>r"
using u w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<sigma>.T1 [of u \<theta>\<^sub>r] by auto
text \<open>
$$
\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}
}
$$
\<close>
have w\<^sub>s: "ide w\<^sub>s"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
have [simp]: "src w\<^sub>s = src u"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s src_cod
by (metis \<sigma>.leg0_simps(2) \<sigma>.tab_simps(2) \<theta>\<^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 \<sigma>"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s
by (metis \<sigma>.tab_simps(2) arr_dom in_homE not_arr_null seq_if_composable)
have \<theta>\<^sub>s_in_hom [intro]: "\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> w\<^sub>s \<Rightarrow> u\<guillemotright>"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
have \<theta>\<^sub>s_in_hhom [intro]: "\<guillemotleft>\<theta>\<^sub>s : src u \<rightarrow> src s\<guillemotright>"
using \<theta>\<^sub>s_in_hom src_cod trg_cod
by (metis \<theta>\<^sub>r_in_hom arrI hseqE in_hhom_def seqE vconn_implies_hpar(1)
vconn_implies_hpar(3) w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s)
have [simp]: "src \<theta>\<^sub>s = src u" using \<theta>\<^sub>s_in_hhom by auto
have [simp]: "trg \<theta>\<^sub>s = src s" using \<theta>\<^sub>s_in_hhom by auto
have [simp]: "dom \<theta>\<^sub>s = s\<^sub>0 \<star> w\<^sub>s" using \<theta>\<^sub>s_in_hom by blast
have [simp]: "cod \<theta>\<^sub>s = u" using \<theta>\<^sub>s_in_hom by blast
have \<nu>\<^sub>s_in_hom [intro]: "\<guillemotleft>\<nu>\<^sub>s : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s\<^sub>1 \<star> w\<^sub>s\<guillemotright>" using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
have \<nu>\<^sub>s_in_hhom [intro]: "\<guillemotleft>\<nu>\<^sub>s : src u \<rightarrow> trg s\<guillemotright>"
using \<nu>\<^sub>s_in_hom src_dom trg_cod
by (metis \<open>src \<theta>\<^sub>r = src u\<close> \<open>trg \<theta>\<^sub>r = trg s\<close> \<theta>\<^sub>r_in_hom in_hhomI in_homE src_dom trg_dom)
have [simp]: "src \<nu>\<^sub>s = src u" using \<nu>\<^sub>s_in_hhom by auto
have [simp]: "trg \<nu>\<^sub>s = trg s" using \<nu>\<^sub>s_in_hhom by auto
have [simp]: "dom \<nu>\<^sub>s = r\<^sub>0 \<star> w\<^sub>r" using \<nu>\<^sub>s_in_hom by auto
have [simp]: "cod \<nu>\<^sub>s = s\<^sub>1 \<star> w\<^sub>s" using \<nu>\<^sub>s_in_hom by auto
have iso_\<nu>\<^sub>s: "iso \<nu>\<^sub>s" using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
obtain w \<theta>\<^sub>t \<nu>\<^sub>t
where w\<theta>\<^sub>t\<nu>\<^sub>t: "ide w \<and> \<guillemotleft>\<theta>\<^sub>t : p\<^sub>0 \<star> w \<Rightarrow> w\<^sub>s\<guillemotright> \<and> \<guillemotleft>\<nu>\<^sub>t : w\<^sub>r \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>\<^sub>t \<and>
(s\<^sub>1 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t) = \<nu>\<^sub>s"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s iso_\<nu>\<^sub>s r\<^sub>0s\<^sub>1.\<phi>_biuniversal_prop(1) [of w\<^sub>s w\<^sub>r \<nu>\<^sub>s] by blast
text \<open>
$$
\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}
}
$$
\<close>
text \<open>{\bf Note:} \<open>w\<close> is not necessarily a map.\<close>
have w: "ide w"
using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
have [simp]: "src w = src u"
using w\<theta>\<^sub>t\<nu>\<^sub>t src_cod
by (metis \<nu>\<^sub>s_in_hom \<open>src \<nu>\<^sub>s = src u\<close> arrI seqE src_hcomp src_vcomp vseq_implies_hpar(1))
have [simp]: "trg w = src p\<^sub>0"
using w\<theta>\<^sub>t\<nu>\<^sub>t
by (metis \<nu>\<^sub>s_in_hom arrI not_arr_null r\<^sub>0s\<^sub>1.\<phi>_simps(2) seqE seq_if_composable)
have \<theta>\<^sub>t_in_hom [intro]: "\<guillemotleft>\<theta>\<^sub>t : p\<^sub>0 \<star> w \<Rightarrow> w\<^sub>s\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
have \<theta>\<^sub>t_in_hhom [intro]: "\<guillemotleft>\<theta>\<^sub>t : src u \<rightarrow> src \<sigma>\<guillemotright>"
using \<theta>\<^sub>t_in_hom src_cod trg_cod \<open>src w\<^sub>s = src u\<close> \<open>trg w\<^sub>s = src \<sigma>\<close> by fastforce
have [simp]: "src \<theta>\<^sub>t = src u" using \<theta>\<^sub>t_in_hhom by auto
have [simp]: "trg \<theta>\<^sub>t = src \<sigma>" using \<theta>\<^sub>t_in_hhom by auto
have [simp]: "dom \<theta>\<^sub>t = p\<^sub>0 \<star> w" using \<theta>\<^sub>t_in_hom by blast
have (* [simp]: *) "cod \<theta>\<^sub>t = w\<^sub>s" using \<theta>\<^sub>t_in_hom by blast
have \<nu>\<^sub>t_in_hom [intro]: "\<guillemotleft>\<nu>\<^sub>t : w\<^sub>r \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright>" using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
have \<nu>\<^sub>t_in_hhom [intro]: "\<guillemotleft>\<nu>\<^sub>t : src u \<rightarrow> src \<rho>\<guillemotright>"
using \<nu>\<^sub>t_in_hom src_dom trg_dom \<open>src w\<^sub>r = src u\<close> \<open>trg w\<^sub>r = src \<rho>\<close> by fastforce
have [simp]: "src \<nu>\<^sub>t = src u" using \<nu>\<^sub>t_in_hhom by auto
have [simp]: "trg \<nu>\<^sub>t = src \<rho>" using \<nu>\<^sub>t_in_hhom by auto
have (* [simp]: *) "dom \<nu>\<^sub>t = w\<^sub>r" using \<nu>\<^sub>t_in_hom by auto
have [simp]: "cod \<nu>\<^sub>t = p\<^sub>1 \<star> w" using \<nu>\<^sub>t_in_hom by auto
have iso_\<nu>\<^sub>t: "iso \<nu>\<^sub>t" using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
define \<theta> where "\<theta> = \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w]"
have \<theta>: "\<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>"
proof (unfold \<theta>_def, intro comp_in_homI)
show "\<guillemotleft>\<a>[s\<^sub>0, p\<^sub>0, w] : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> s\<^sub>0 \<star> p\<^sub>0 \<star> w\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by auto
show "\<guillemotleft>s\<^sub>0 \<star> \<theta>\<^sub>t : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> s\<^sub>0 \<star> w\<^sub>s\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by auto
show "\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> w\<^sub>s \<Rightarrow> u\<guillemotright>"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
qed
define \<nu> where "\<nu> = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
have \<nu>: "\<guillemotleft>\<nu> : ?v \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright>"
proof (unfold \<nu>_def, intro comp_in_homI)
show "\<guillemotleft>\<nu>\<^sub>r : ?v \<Rightarrow> r\<^sub>1 \<star> w\<^sub>r\<guillemotright>"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by blast
show "\<guillemotleft>r\<^sub>1 \<star> \<nu>\<^sub>t : r\<^sub>1 \<star> w\<^sub>r \<Rightarrow> r\<^sub>1 \<star> p\<^sub>1 \<star> w\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] : r\<^sub>1 \<star> p\<^sub>1 \<star> w \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t assoc_in_hom by (simp add: \<rho>.T0.antipar(1))
qed
have iso_\<nu>: "iso \<nu>"
using \<nu> w\<theta>\<^sub>t\<nu>\<^sub>t w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<rho>.T0.antipar(1)
by (unfold \<nu>_def, intro isos_compose) auto
have *: "arr ((s \<star> \<theta>\<^sub>s) \<cdot> \<a>[s, s\<^sub>0, w\<^sub>s] \<cdot> (\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s w\<theta>\<^sub>t\<nu>\<^sub>t \<theta>\<^sub>r_in_hom comp_assoc by auto
have "((r \<star> s) \<star> \<theta>) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu> = \<omega>"
proof -
have "seq (r \<star> \<theta>\<^sub>r) (\<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r)"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<rho>.base_simps(2) composable by fastforce
hence "\<omega> = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> \<rho>.composite_cell w\<^sub>r \<theta>\<^sub>r \<cdot> \<nu>\<^sub>r"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r invert_side_of_triangle(1) iso_assoc
by (metis 1 \<rho>.ide_base \<sigma>.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\<theta>\<^sub>s\<nu>\<^sub>s)
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> \<rho>.composite_cell w\<^sub>r (\<sigma>.composite_cell w\<^sub>s \<theta>\<^sub>s \<cdot> \<nu>\<^sub>s) \<cdot> \<nu>\<^sub>r"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> \<a>[s, s\<^sub>0, w\<^sub>s] \<cdot>
(\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
using w\<theta>\<^sub>t\<nu>\<^sub>t comp_assoc by simp
text \<open>Rearrange to create \<open>\<theta>\<close> and \<open>\<nu>\<close>, leaving \<open>tab\<close> in the middle.\<close>
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> \<a>[s, s\<^sub>0, w\<^sub>s] \<cdot>
((\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t)) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
using comp_assoc by presburger
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (\<a>[s, s\<^sub>0, w\<^sub>s] \<cdot>
((s \<star> s\<^sub>0) \<star> \<theta>\<^sub>t)) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
proof -
have "(\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) = \<sigma> \<star> \<theta>\<^sub>t"
using comp_arr_dom comp_cod_arr interchange
by (metis \<open>cod \<theta>\<^sub>t = w\<^sub>s\<close> \<sigma>.tab_simps(1) \<sigma>.tab_simps(4) arrI w\<theta>\<^sub>t\<nu>\<^sub>t)
also have "... = ((s \<star> s\<^sub>0) \<star> \<theta>\<^sub>t) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
using comp_arr_dom comp_cod_arr interchange w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s w\<theta>\<^sub>t\<nu>\<^sub>t \<sigma>.tab_in_hom
by (metis \<open>dom \<theta>\<^sub>t = p\<^sub>0 \<star> w\<close> \<sigma>.tab_simps(5) arrI)
finally have "(\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) = ((s \<star> s\<^sub>0) \<star> \<theta>\<^sub>t) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
using assoc_naturality [of s s\<^sub>0 \<theta>\<^sub>t] w\<theta>\<^sub>t\<nu>\<^sub>t comp_assoc \<open>cod \<theta>\<^sub>t = w\<^sub>s\<close> arrI by force
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot>
(\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
proof -
have "r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t) =
(r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
proof -
have "seq ((s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t))
(\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
proof -
have "seq (s \<star> \<theta>\<^sub>s)
((s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
using \<open>\<guillemotleft>\<a>[r, s, u] \<cdot> \<omega> : dom \<omega> \<Rightarrow> r \<star> s \<star> u\<guillemotright>\<close> calculation by blast
thus ?thesis
using comp_assoc by presburger
qed
thus ?thesis
using whisker_left [of r "(s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
"\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)"]
w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s w\<theta>\<^sub>t\<nu>\<^sub>t comp_assoc
by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> ((r \<star> r\<^sub>0 \<star> \<nu>\<^sub>t) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r]) \<cdot>
(\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
proof -
have "seq (\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) (r\<^sub>0 \<star> \<nu>\<^sub>t)"
using 1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\<theta>\<^sub>t\<nu>\<^sub>t
apply (intro seqI' comp_in_homI) by auto
hence "r \<star> (\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t) =
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> (r \<star> r\<^sub>0 \<star> \<nu>\<^sub>t)"
using whisker_left by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot>
(((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t) \<cdot> (\<rho> \<star> w\<^sub>r)) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> r\<^sub>0 \<star> \<nu>\<^sub>t) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] = \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> ((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t)"
using assoc_naturality [of r r\<^sub>0 \<nu>\<^sub>t] \<nu>\<^sub>t_in_hom by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t))) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t) \<cdot> (\<rho> \<star> w\<^sub>r) = \<rho> \<star> \<nu>\<^sub>t"
using comp_arr_dom comp_cod_arr interchange
by (metis \<open>dom \<nu>\<^sub>t = w\<^sub>r\<close> \<rho>.tab_simps(1) \<rho>.tab_simps(5) arrI w\<theta>\<^sub>t\<nu>\<^sub>t)
also have "... = (\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t)"
using comp_arr_dom comp_cod_arr interchange
by (metis \<open>cod \<nu>\<^sub>t = p\<^sub>1 \<star> w\<close> \<open>trg \<nu>\<^sub>t = src \<rho>\<close> \<rho>.T0.antipar(1) \<rho>.tab_simps(1)
\<rho>.tab_simps(2) \<rho>.tab_simps(4) r\<^sub>0s\<^sub>1.base_simps(2) trg.preserves_reflects_arr
trg_hcomp)
finally have "((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t) \<cdot> (\<rho> \<star> w\<^sub>r) = (\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t)" by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> ((\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) =
((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
proof -
have "seq (s \<star> \<theta>\<^sub>s) (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
using \<theta>\<^sub>s_in_hom \<theta>\<^sub>s_in_hhom \<theta>\<^sub>t_in_hom \<theta>\<^sub>t_in_hhom 1 calculation by blast
moreover have "src r = trg (s \<star> \<theta>\<^sub>s)"
using composable hseqI by force
ultimately
have "\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) =
(\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> s \<star> \<theta>\<^sub>s)) \<cdot> (r \<star> s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
using whisker_left comp_assoc by simp
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> w\<^sub>s] \<cdot> (r \<star> s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
using assoc_naturality comp_assoc
by (metis \<open>cod \<theta>\<^sub>s = u\<close> \<open>dom \<theta>\<^sub>s = s\<^sub>0 \<star> w\<^sub>s\<close> \<open>trg \<theta>\<^sub>s = src s\<close>
\<rho>.base_simps(2-4) \<sigma>.base_simps(2-4) arrI assoc'_naturality composable w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s)
also have "... = (((r \<star> s) \<star> \<theta>\<^sub>s) \<cdot> ((r \<star> s) \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> w\<^sub>s] \<cdot> (r \<star> s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) =
((r \<star> s) \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
using arrI hseq_char assoc'_naturality [of r s "s\<^sub>0 \<star> \<theta>\<^sub>t"] \<open>cod \<theta>\<^sub>t = w\<^sub>s\<close> composable
by auto
thus ?thesis
using comp_assoc by auto
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
using \<theta>_def \<theta> 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 \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot>
((r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w] \<cdot>
((\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] = \<a>[s \<star> s\<^sub>0, p\<^sub>0, w] \<cdot> ((\<sigma> \<star> p\<^sub>0) \<star> w)"
using assoc_naturality [of \<sigma> p\<^sub>0 w] by (simp add: w\<theta>\<^sub>t\<nu>\<^sub>t)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]) \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
using r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\<theta>\<^sub>t\<nu>\<^sub>t whisker_left comp_assoc by force
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
(\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]) \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w)) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] =
\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]"
proof -
have 1: "(r \<star> \<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) =
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"
using pentagon
by (simp add: \<rho>.T0.antipar(1) w)
moreover have 2: "seq \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w by simp
moreover have "inv (r \<star> \<a>[r\<^sub>0, p\<^sub>1, w]) = r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w by simp
ultimately have "\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) =
((r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]) \<cdot> \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w comp_assoc
invert_side_of_triangle(1)
[of "\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]" "r \<star> \<a>[r\<^sub>0, p\<^sub>1, w]"
"\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w)"]
by simp
hence "(r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] =
(\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w
invert_side_of_triangle(2)
[of "\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w)"
"(r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]" "\<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"]
using \<open>trg w = src p\<^sub>0\<close> by simp
thus ?thesis
using comp_assoc by presburger
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w]) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w) = ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using assoc'_naturality [of \<rho> p\<^sub>1 w] by (simp add: \<rho>.T0.antipar(1) w\<theta>\<^sub>t\<nu>\<^sub>t)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> ((r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>[r, s\<^sub>1 \<star> p\<^sub>0, w]) \<cdot>
((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] = \<a>[r, s\<^sub>1 \<star> p\<^sub>0, w] \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w)"
using assoc_naturality [of r r\<^sub>0s\<^sub>1.\<phi> w] r\<^sub>0s\<^sub>1.cospan w\<theta>\<^sub>t\<nu>\<^sub>t by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot>
(((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot>
((\<rho> \<star> p\<^sub>1) \<star> w)) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>[r, s\<^sub>1 \<star> p\<^sub>0, w] =
\<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> ((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w)"
proof -
have "arr w \<and> dom w = w \<and> cod w = w"
using ide_char w by blast
then show ?thesis
using assoc_naturality [of r "\<sigma> \<star> p\<^sub>0" w] composable by auto
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w)) \<cdot>
(tab \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) =
(r \<star> \<sigma> \<star> p\<^sub>0) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1) \<star> w"
using w \<rho>.T0.antipar(1) composable whisker_right by auto
also have "... = (((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1) \<star> w"
proof -
have "((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) =
r \<star> \<sigma> \<star> p\<^sub>0"
proof -
have "((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) =
((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> ((r \<star> s \<star> s\<^sub>0 \<star> p\<^sub>0) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]))) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)"
using comp_assoc_assoc' by (simp add: composable)
also have "... = ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)"
using comp_cod_arr by (simp add: composable)
also have "... = ((r \<star> (s \<star> s\<^sub>0) \<star> p\<^sub>0)) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)"
using whisker_left comp_assoc_assoc' assoc_in_hom hseqI'
by (metis \<rho>.ide_base \<sigma>.base_simps(2) \<sigma>.ide_base \<sigma>.ide_leg0
\<sigma>.leg0_simps(2-3) \<sigma>.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 \<star> \<sigma> \<star> 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 \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> \<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1) \<star> w"
using comp_assoc by presburger
also have "... = (r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> \<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> tab \<star> w"
using tab_def by simp
also have "... = ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) \<cdot> (tab \<star> w)"
using w \<rho>.T0.antipar(1) composable comp_assoc whisker_right by auto
finally have "((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot>
((\<rho> \<star> p\<^sub>1) \<star> w) =
((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) \<cdot> (tab \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> ((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w])) \<cdot>
\<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot> (r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
\<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) =
((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot> (r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
\<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot>
(\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, (\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>) \<^bold>\<cdot> (\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>)\<rbrace>"
using w comp_assoc \<a>'_def \<alpha>_def composable by simp
also have "... = \<lbrace>((\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using w composable
apply (intro E.eval_eqI) by simp_all
also have "... = ((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w]"
using w comp_assoc \<a>'_def \<alpha>_def composable by simp
finally show ?thesis by simp
qed
thus ?thesis
using \<nu>_def comp_assoc by simp
qed
also have "... = ((r \<star> s) \<star> \<theta>) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>"
proof -
have "((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> ((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w]) = (r \<star> s) \<star> \<theta>"
using \<theta>_def w whisker_left composable
by (metis \<theta> arrI ide_base comp_assoc)
thus ?thesis
using comp_assoc by presburger
qed
finally show "((r \<star> s) \<star> \<theta>) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu> = \<omega>"
by simp
qed
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
using w\<theta>\<^sub>t\<nu>\<^sub>t \<theta> \<nu> iso_\<nu> comp_assoc by metis
qed
show "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w'; \<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : (s\<^sub>0 \<star> p\<^sub>0) \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : (r\<^sub>1 \<star> p\<^sub>1) \<star> w \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w'\<guillemotright>;
composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume w: "ide w"
assume w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : (s\<^sub>0 \<star> p\<^sub>0) \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : (r\<^sub>1 \<star> p\<^sub>1) \<star> w \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w'\<guillemotright>"
assume eq: "composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta>"
interpret uw\<theta>w'\<theta>'\<beta>: uw\<theta>w'\<theta>'\<beta> V H \<a> \<i> src trg
\<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close> u w \<theta> w' \<theta>' \<beta>
using w w' \<theta> \<theta>' \<beta> eq composable tab_in_hom comp_assoc
by unfold_locales auto
text \<open>
$$
\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}
$$
\<close>
text \<open>
First apply property \<open>\<rho>.T2\<close> using \<open>\<guillemotleft>\<beta>\<^sub>r : r\<^sub>1 \<star> p\<^sub>1 \<star> w \<Rightarrow> r\<^sub>1 \<star> p\<^sub>1 \<star> w'\<guillemotright>\<close>
(obtained by composing \<open>\<beta>\<close> with associativities) and ``everything to the right''
as \<open>\<theta>\<^sub>r\<close> and \<open>\<theta>\<^sub>r'\<close>. This yields \<open>\<guillemotleft>\<gamma>\<^sub>r : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>\<close>.
Next, apply property \<open>\<rho>.T2\<close> to obtain \<open>\<guillemotleft>\<gamma>\<^sub>s : p\<^sub>0 \<star> w \<Rightarrow> p\<^sub>0 \<star> w'\<guillemotright>\<close>.
For this use \<open>\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>\<close> and \<open>\<guillemotleft>\<theta>\<^sub>s' : s\<^sub>0 \<star> p\<^sub>0 \<star> w'\<guillemotright>\<close>
obtained by composing \<open>\<theta>\<close> and \<open>\<theta>'\<close> with associativities.
We also need \<open>\<guillemotleft>\<beta>\<^sub>s : s\<^sub>1 \<star> p\<^sub>0 \<star> w \<Rightarrow> s\<^sub>1 \<star> p\<^sub>0 \<star> w'\<guillemotright>\<close>.
To get this, transport \<open>r\<^sub>0 \<star> \<gamma>\<^sub>r\<close> across \<open>\<phi>\<close>; we need \<open>\<phi>\<close> to be an isomorphism
in order to do this.
Finally, apply the biuniversal property of \<open>\<phi>\<close> to get \<open>\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>\<close>
and verify the required equation.
\<close>
let ?w\<^sub>r = "p\<^sub>1 \<star> w"
have w\<^sub>r: "ide ?w\<^sub>r" by simp
let ?w\<^sub>r' = "p\<^sub>1 \<star> w'"
have w\<^sub>r': "ide ?w\<^sub>r'" by simp
define \<theta>\<^sub>r where "\<theta>\<^sub>r = (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
have \<theta>\<^sub>r: "\<guillemotleft>\<theta>\<^sub>r : r\<^sub>0 \<star> ?w\<^sub>r \<Rightarrow> s \<star> u\<guillemotright>"
unfolding \<theta>\<^sub>r_def
using \<rho>.T0.antipar(1) by fastforce
define \<theta>\<^sub>r' where "\<theta>\<^sub>r' = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']"
have \<theta>\<^sub>r': "\<guillemotleft>\<theta>\<^sub>r' : r\<^sub>0 \<star> ?w\<^sub>r' \<Rightarrow> s \<star> u\<guillemotright>"
unfolding \<theta>\<^sub>r'_def
using \<rho>.T0.antipar(1) by fastforce
define \<beta>\<^sub>r where "\<beta>\<^sub>r = \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
have \<beta>\<^sub>r: "\<guillemotleft>\<beta>\<^sub>r : r\<^sub>1 \<star> ?w\<^sub>r \<Rightarrow> r\<^sub>1 \<star> ?w\<^sub>r'\<guillemotright>"
unfolding \<beta>\<^sub>r_def
using \<rho>.T0.antipar(1) by force
have eq\<^sub>r: "\<rho>.composite_cell ?w\<^sub>r \<theta>\<^sub>r = \<rho>.composite_cell ?w\<^sub>r' \<theta>\<^sub>r' \<cdot> \<beta>\<^sub>r"
text \<open>
$$
\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}
$$
\<close>
proof -
have "\<rho>.composite_cell ?w\<^sub>r \<theta>\<^sub>r = \<a>[r, s, u] \<cdot> composite_cell w \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using \<theta>\<^sub>r_def technical uw\<theta>w'\<theta>'\<beta>.uw\<theta>.uw\<theta> by blast
also have "... = \<a>[r, s, u] \<cdot> (((r \<star> s) \<star> \<theta>') \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w'] \<cdot>
(tab \<star> w') \<cdot> \<beta>) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using eq comp_assoc by simp
also have "... = (r \<star> \<theta>\<^sub>r') \<cdot> \<a>[r, r\<^sub>0, ?w\<^sub>r'] \<cdot> (\<rho> \<star> ?w\<^sub>r') \<cdot> \<beta>\<^sub>r"
proof -
have "\<a>[r, s, u] \<cdot> (composite_cell w' \<theta>' \<cdot> \<beta>) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] =
\<a>[r, s, u] \<cdot> composite_cell w' \<theta>' \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = (r \<star> (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w'] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w') \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
proof -
have "\<a>[r, s, u] \<cdot> composite_cell w' \<theta>' \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] =
\<a>[r, s, u] \<cdot> composite_cell w' \<theta>' \<cdot>
((\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<beta>) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
proof -
have "(\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<beta> = \<beta>"
using comp_cod_arr \<rho>.T0.antipar(1) \<beta> comp_assoc_assoc' by simp
thus ?thesis by argo
qed
also have "... = (\<a>[r, s, u] \<cdot> ((r \<star> s) \<star> \<theta>') \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w'] \<cdot> (tab \<star> w') \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = ((r \<star> (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w'] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w')) \<cdot>
\<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using \<theta>\<^sub>r'_def technical [of w' \<theta>' u ?w\<^sub>r' \<theta>\<^sub>r'] comp_assoc by fastforce
finally show ?thesis
using comp_assoc by simp
qed
finally show ?thesis
using \<theta>\<^sub>r'_def \<beta>\<^sub>r_def comp_assoc by auto
qed
finally show ?thesis
using comp_assoc by presburger
qed
have 1: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright> \<and> \<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>r = r\<^sub>1 \<star> \<gamma>"
using eq\<^sub>r \<rho>.T2 [of ?w\<^sub>r ?w\<^sub>r' \<theta>\<^sub>r "s \<star> u" \<theta>\<^sub>r' \<beta>\<^sub>r] w\<^sub>r w\<^sub>r' \<theta>\<^sub>r \<theta>\<^sub>r' \<beta>\<^sub>r by blast
obtain \<gamma>\<^sub>r where \<gamma>\<^sub>r: "\<guillemotleft>\<gamma>\<^sub>r : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright> \<and> \<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<and> \<beta>\<^sub>r = r\<^sub>1 \<star> \<gamma>\<^sub>r"
using 1 by blast
let ?w\<^sub>s = "p\<^sub>0 \<star> w"
have w\<^sub>s: "ide ?w\<^sub>s" by simp
let ?w\<^sub>s' = "p\<^sub>0 \<star> w'"
have w\<^sub>s': "ide ?w\<^sub>s'" by simp
define \<theta>\<^sub>s where "\<theta>\<^sub>s = \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
have \<theta>\<^sub>s: "\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>"
using \<theta>\<^sub>s_def by auto
define \<theta>\<^sub>s' where "\<theta>\<^sub>s' = \<theta>' \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']"
have \<theta>\<^sub>s': "\<guillemotleft>\<theta>\<^sub>s' : s\<^sub>0 \<star> p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>"
using \<theta>\<^sub>s'_def by auto
define \<beta>\<^sub>s where "\<beta>\<^sub>s = \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
have \<beta>\<^sub>s: "\<guillemotleft>\<beta>\<^sub>s : s\<^sub>1 \<star> ?w\<^sub>s \<Rightarrow> s\<^sub>1 \<star> ?w\<^sub>s'\<guillemotright>"
unfolding \<beta>\<^sub>s_def
using \<gamma>\<^sub>r r\<^sub>0s\<^sub>1.\<phi>_in_hom(2) r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) \<rho>.T0.antipar(1)
apply (intro comp_in_homI)
apply auto
by auto
have eq\<^sub>s: "\<sigma>.composite_cell (p\<^sub>0 \<star> w) (\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) =
\<sigma>.composite_cell (p\<^sub>0 \<star> w') (\<theta>' \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<beta>\<^sub>s"
text \<open>
$$
\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}
$$
\<close>
proof -
have "\<sigma>.composite_cell (p\<^sub>0 \<star> w') (\<theta>' \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<beta>\<^sub>s =
(\<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using \<beta>\<^sub>s_def \<theta>\<^sub>r'_def whisker_left comp_assoc by simp
also have "... = \<theta>\<^sub>r \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using \<gamma>\<^sub>r by simp
also have "... = ((s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \<cdot> \<a>[s, s\<^sub>0, ?w\<^sub>s] \<cdot> (\<sigma> \<star> ?w\<^sub>s) \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using \<theta>\<^sub>r_def comp_assoc by simp
also have "... = (s \<star> \<theta>) \<cdot> \<sigma>.composite_cell (p\<^sub>0 \<star> w) \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
proof -
have "(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] =
\<sigma> \<star> p\<^sub>0 \<star> w"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) \<rho>.T0.antipar(1) comp_assoc_assoc' by simp
text \<open>Here the fact that \<open>\<phi>\<close> is a retraction is used.\<close>
moreover have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) = cod \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_arr_inv' whisker_right [of w r\<^sub>0s\<^sub>1.\<phi> "inv r\<^sub>0s\<^sub>1.\<phi>"]
by simp
moreover have "\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = dom (\<sigma> \<star> p\<^sub>0 \<star> 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.\<phi>) w"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2)
by (intro hseqI, auto)
moreover have "hseq \<sigma> (p\<^sub>0 \<star> 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 "... = \<sigma>.composite_cell (p\<^sub>0 \<star> w) (\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])"
using \<theta>\<^sub>s_def whisker_left
by (metis \<sigma>.ide_base \<theta>\<^sub>s arrI comp_assoc)
finally show ?thesis by simp
qed
hence 2: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright> \<and> \<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>s = s\<^sub>1 \<star> \<gamma>"
using \<sigma>.T2 [of ?w\<^sub>s ?w\<^sub>s' \<theta>\<^sub>s u \<theta>\<^sub>s' \<beta>\<^sub>s] w\<^sub>s w\<^sub>s' \<theta>\<^sub>s \<theta>\<^sub>s' \<beta>\<^sub>s
by (metis \<theta>\<^sub>s'_def \<theta>\<^sub>s_def)
obtain \<gamma>\<^sub>s where \<gamma>\<^sub>s: "\<guillemotleft>\<gamma>\<^sub>s : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright> \<and> \<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> \<gamma>\<^sub>s) \<and> \<beta>\<^sub>s = s\<^sub>1 \<star> \<gamma>\<^sub>s"
using 2 by blast
have eq\<^sub>t: "(s\<^sub>1 \<star> \<gamma>\<^sub>s) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
(s\<^sub>1 \<star> ?w\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
text \<open>
$$
\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
$$
\<close>
proof -
have "(s\<^sub>1 \<star> ?w\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) =
\<beta>\<^sub>s \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "\<beta>\<^sub>s \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
(\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<beta>\<^sub>s_def comp_assoc by metis
also have "... = (\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
proof -
have "(r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
r\<^sub>0 \<star> \<gamma>\<^sub>r"
proof -
have "(r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
(r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_assoc_assoc' comp_cod_arr by simp
(* Used here that \<phi> is a section. *)
also have "... = (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_inv_arr' \<rho>.T0.antipar(1)
whisker_right [of w "inv r\<^sub>0s\<^sub>1.\<phi>" r\<^sub>0s\<^sub>1.\<phi>] comp_cod_arr
by simp
also have "... = r\<^sub>0 \<star> \<gamma>\<^sub>r"
proof -
have "hseq r\<^sub>0 \<gamma>\<^sub>r"
using \<beta>\<^sub>s \<beta>\<^sub>s_def by blast
thus ?thesis
using comp_assoc_assoc' comp_arr_dom
by (metis (no_types) \<gamma>\<^sub>r \<rho>.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 "... = \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
using comp_assoc by presburger
also have "... = (s\<^sub>1 \<star> ?w\<^sub>s') \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
proof -
have "(s\<^sub>1 \<star> ?w\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] = \<a>[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 \<star> \<gamma>\<^sub>s) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>\<^sub>s by simp
finally show ?thesis by simp
qed
have 3: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<gamma>\<^sub>s = (p\<^sub>0 \<star> w') \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<gamma>\<^sub>r"
using w w' w\<^sub>s' w\<^sub>r \<gamma>\<^sub>r \<gamma>\<^sub>s eq\<^sub>t
r\<^sub>0s\<^sub>1.\<phi>_biuniversal_prop(2) [of ?w\<^sub>s' ?w\<^sub>r w w' \<gamma>\<^sub>s "p\<^sub>0 \<star> w'" \<gamma>\<^sub>r]
by blast
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<gamma>\<^sub>s = (p\<^sub>0 \<star> w') \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<gamma>\<^sub>r"
using 3 by blast
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
have "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
have "\<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
have "\<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>) = (\<theta>\<^sub>s' \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w']) \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
using \<theta>\<^sub>s'_def comp_arr_dom comp_assoc comp_assoc_assoc'(2) by auto
also have "... = (\<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>)) \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w]"
using assoc_naturality [of s\<^sub>0 p\<^sub>0 \<gamma>] comp_assoc
by (metis \<gamma> \<gamma>\<^sub>r \<sigma>.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 "... = \<theta>\<^sub>s \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w]"
by (metis \<gamma> \<gamma>\<^sub>s arrI comp_ide_arr w\<^sub>s')
also have "... = \<theta>"
using \<theta>\<^sub>s_def comp_assoc comp_arr_dom comp_assoc_assoc' by simp
finally show ?thesis by simp
qed
moreover have "\<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>"
proof -
have "\<beta> = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta>\<^sub>r \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta>\<^sub>r \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w] =
(\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w]"
using \<beta>\<^sub>r_def comp_assoc by simp
also have "... = \<beta>"
using comp_arr_dom comp_cod_arr
by (metis \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 comp_assoc_assoc'(2) hseqE ideD(1)
uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(4-5) leg1_simps(2) w w' w\<^sub>r w\<^sub>r')
finally show ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> (r\<^sub>1 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w]"
using \<gamma>\<^sub>r by simp
also have "... = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> ((r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>)"
using assoc_naturality [of r\<^sub>1 p\<^sub>1 \<gamma>]
by (metis \<gamma> \<gamma>\<^sub>r \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE
ide_char in_homE leg1_simps(2))
also have "... = (\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> ((r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>)"
using comp_assoc by presburger
also have "... = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>"
using comp_cod_arr
by (metis \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 calculation comp_assoc_assoc'(2) comp_ide_arr
hseqE ideD(1) ide_cod local.uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) local.uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(5)
w' w\<^sub>r')
finally show ?thesis by simp
qed
ultimately show "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
using \<gamma> by blast
qed
moreover have "\<And>\<gamma>'. \<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>')
\<Longrightarrow> \<gamma>' = \<gamma>"
proof -
fix \<gamma>'
assume \<gamma>': "\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>')"
show "\<gamma>' = \<gamma>"
proof -
let ?P\<^sub>r = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright> \<and> \<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>r = r\<^sub>1 \<star> \<gamma>"
let ?P\<^sub>s = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright> \<and> \<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>s = s\<^sub>1 \<star> \<gamma>"
let ?\<gamma>\<^sub>r' = "p\<^sub>1 \<star> \<gamma>'"
let ?\<gamma>\<^sub>s' = "p\<^sub>0 \<star> \<gamma>'"
let ?P\<^sub>t = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<gamma>\<^sub>s = (p\<^sub>0 \<star> w') \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<gamma>\<^sub>r"
have "hseq p\<^sub>0 \<gamma>'"
proof (intro hseqI)
show "\<guillemotleft>\<gamma>' : src \<theta> \<rightarrow> src p\<^sub>0\<guillemotright>"
using \<gamma>'
by (metis hseqE hseqI' in_hhom_def uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) src_hcomp
src_vcomp leg0_simps(2) leg1_simps(3)
uw\<theta>w'\<theta>'\<beta>.uw\<theta>.\<theta>_simps(1) vseq_implies_hpar(1))
show "\<guillemotleft>p\<^sub>0 : src p\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright>" by simp
qed
hence "hseq p\<^sub>1 \<gamma>'"
using hseq_char by simp
have "\<guillemotleft>?\<gamma>\<^sub>r' : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright>"
using \<gamma>' by auto
moreover have "\<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r')"
proof -
text \<open>
Note that @{term \<theta>\<^sub>r} is the composite of ``everything to the right''
of @{term "\<rho> \<star> ?w\<^sub>r"}, and similarly for @{term \<theta>\<^sub>r'}.
We can factor @{term \<theta>\<^sub>r} as @{term "(s \<star> \<theta>) \<cdot> X w"}, where @{term "X w"}
is a composite of @{term \<sigma>} and @{term \<phi>}. We can similarly factor @{term \<theta>\<^sub>r'}
as @{term "(s \<star> \<theta>') \<cdot> X w'"}.
Then @{term "\<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r') = (s \<star> \<theta>') \<cdot> X w' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r')"},
which equals @{term "(s \<star> \<theta>') \<cdot> (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> ?\<gamma>\<^sub>r') \<cdot> X w = \<theta>\<^sub>r"}.
\<close>
let ?X = "\<lambda>w. (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
have "\<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r') = (s \<star> \<theta>') \<cdot> ?X w' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r')"
using \<theta>\<^sub>r'_def comp_assoc by simp
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> ?X w"
proof -
have "(s \<star> \<theta>') \<cdot> ((s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>') =
(s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>')"
using comp_assoc by presburger
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot>
((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>')) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using assoc'_naturality [of r\<^sub>0 p\<^sub>1 \<gamma>'] comp_assoc
by (metis \<gamma>' \<open>\<guillemotleft>p\<^sub>1 \<star> \<gamma>' : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>\<close> \<rho>.T0.antipar(1)
\<rho>.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 \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>')) \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') = r\<^sub>0s\<^sub>1.\<phi> \<star> \<gamma>'"
using \<gamma>' interchange [of r\<^sub>0s\<^sub>1.\<phi> "r\<^sub>0 \<star> p\<^sub>1" w' \<gamma>'] comp_arr_dom comp_cod_arr
by auto
also have "... = ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using \<gamma>' interchange \<open>hseq p\<^sub>0 \<gamma>'\<close> comp_arr_dom comp_cod_arr
by (metis comp_arr_ide r\<^sub>0s\<^sub>1.\<phi>_simps(1,5) seqI'
uw\<theta>w'\<theta>'\<beta>.uw\<theta>.w_in_hom(2) w)
finally have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') =
((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
((\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (s\<^sub>1 \<star> p\<^sub>0 \<star> \<gamma>')) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>' assoc_naturality [of s\<^sub>1 p\<^sub>0 \<gamma>'] comp_assoc
by (metis \<sigma>.leg1_simps(2) \<sigma>.leg1_simps(3,5-6) r\<^sub>0s\<^sub>1.leg0_simps(4-5)
hcomp_in_vhomE hseqE in_homE uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1)
leg0_in_hom(2) leg1_simps(3))
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
((s \<star> s\<^sub>0) \<star> p\<^sub>0 \<star> \<gamma>')) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (s\<^sub>1 \<star> p\<^sub>0 \<star> \<gamma>') = \<sigma> \<star> p\<^sub>0 \<star> \<gamma>'"
using \<gamma>' interchange [of \<sigma> s\<^sub>1 "p\<^sub>0 \<star> w'" "p\<^sub>0 \<star> \<gamma>'"]
whisker_left \<open>hseq p\<^sub>0 \<gamma>'\<close>comp_arr_dom comp_cod_arr
by (metis \<sigma>.tab_simps(1) \<sigma>.tab_simps(4) hcomp_simps(4) in_homE
r\<^sub>0s\<^sub>1.leg0_simps(5))
also have "... = ((s \<star> s\<^sub>0) \<star> p\<^sub>0 \<star> \<gamma>') \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
using \<gamma>' interchange [of "s \<star> s\<^sub>0" \<sigma> "p\<^sub>0 \<star> \<gamma>'" "p\<^sub>0 \<star> w"]
whisker_left comp_arr_dom comp_cod_arr \<open>hseq p\<^sub>0 \<gamma>'\<close>
by (metis \<sigma>.tab_simps(1) \<sigma>.tab_simps(5) hcomp_simps(3) in_homE
r\<^sub>0s\<^sub>1.leg0_simps(4))
finally have "(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (s\<^sub>1 \<star> p\<^sub>0 \<star> \<gamma>') =
((s \<star> s\<^sub>0) \<star> p\<^sub>0 \<star> \<gamma>') \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> ((s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>') \<cdot>
\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>' assoc_naturality [of s s\<^sub>0 "p\<^sub>0 \<star> \<gamma>'"] \<open>hseq p\<^sub>0 \<gamma>'\<close> by force
also have "... = (s \<star> \<theta>') \<cdot> ((s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>')) \<cdot>
\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = (s \<star> \<theta>') \<cdot> ((s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \<cdot>
\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>') =
(s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])"
proof -
have "(s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>') =
s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] \<cdot> (s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>')"
proof -
have "seq \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] (s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>')"
proof
(* It seems to be too time-consuming for auto to solve these. *)
show "\<guillemotleft>s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>' : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> s\<^sub>0 \<star> p\<^sub>0 \<star> w'\<guillemotright>"
using \<gamma>'
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] : s\<^sub>0 \<star> p\<^sub>0 \<star> w' \<Rightarrow> (s\<^sub>0 \<star> p\<^sub>0) \<star> w'\<guillemotright>"
by auto
qed
thus ?thesis
using w w' \<gamma>' whisker_left by simp
qed
also have "... = s \<star> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
using \<gamma>' \<open>hseq p\<^sub>0 \<gamma>'\<close> assoc'_naturality [of s\<^sub>0 p\<^sub>0 \<gamma>'] by fastforce
also have "... = (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])"
proof -
have "seq ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
proof
(* Same here. *)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w] : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> (s\<^sub>0 \<star> p\<^sub>0) \<star> w\<guillemotright>"
by auto
show "\<guillemotleft>(s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>' : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> (s\<^sub>0 \<star> p\<^sub>0) \<star> w'\<guillemotright>"
using \<gamma>' by (intro hcomp_in_vhom, auto)
qed
thus ?thesis
using w w' \<gamma>' whisker_left by simp
qed
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> ?X w"
using comp_assoc by presburger
finally show ?thesis by simp
qed
also have "... = \<theta>\<^sub>r"
using \<theta>\<^sub>r_def \<gamma>' uw\<theta>w'\<theta>'\<beta>.uw\<theta>.\<theta>_simps(1) whisker_left \<sigma>.ide_base comp_assoc
by simp
finally show ?thesis by simp
qed
moreover have "\<beta>\<^sub>r = r\<^sub>1 \<star> ?\<gamma>\<^sub>r'"
proof -
have "\<beta>\<^sub>r = \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> ((r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using \<beta>\<^sub>r_def \<gamma>' by simp
also have "... = \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> (r\<^sub>1 \<star> p\<^sub>1 \<star> \<gamma>')"
using \<gamma>' assoc'_naturality
by (metis \<rho>.leg1_simps(5-6) r\<^sub>0s\<^sub>1.leg1_simps(5-6)
hcomp_in_vhomE hseqE in_homE uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) leg1_in_hom(2))
also have "... = (\<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \<cdot> (r\<^sub>1 \<star> p\<^sub>1 \<star> \<gamma>')"
using comp_assoc by presburger
also have "... = r\<^sub>1 \<star> p\<^sub>1 \<star> \<gamma>'"
using comp_cod_arr
by (metis (no_types, lifting) \<beta>\<^sub>r \<rho>.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 ?\<gamma>\<^sub>r'"
by simp
have eq\<^sub>r: "\<gamma>\<^sub>r = ?\<gamma>\<^sub>r'"
using 1 \<gamma>\<^sub>r P\<^sub>r' by blast
have "\<guillemotleft>?\<gamma>\<^sub>s' : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright>"
using \<gamma>' by auto
moreover have "\<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> ?\<gamma>\<^sub>s')"
using \<gamma>' \<open>hseq p\<^sub>0 \<gamma>'\<close> \<sigma>.leg0_simps(2,4-5) \<sigma>.leg1_simps(3) \<theta>\<^sub>s'_def \<theta>\<^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 "\<beta>\<^sub>s = s\<^sub>1 \<star> ?\<gamma>\<^sub>s'"
proof -
have "\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] =
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>')) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using eq\<^sub>r comp_assoc by simp
also have "... = \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>')) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>') = ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>' assoc'_naturality \<open>hseq p\<^sub>1 \<gamma>'\<close>
by (metis \<rho>.leg0_simps(2,4-5) \<rho>.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 "... = (\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>')) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
proof -
have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') = r\<^sub>0s\<^sub>1.\<phi> \<star> \<gamma>'"
using \<gamma>' interchange [of r\<^sub>0s\<^sub>1.\<phi> "r\<^sub>0 \<star> p\<^sub>1" w' \<gamma>']
comp_arr_dom comp_cod_arr
by auto
also have "... = ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using \<gamma>' interchange \<open>hseq p\<^sub>0 \<gamma>'\<close> comp_arr_dom comp_cod_arr
by (metis in_homE r\<^sub>0s\<^sub>1.\<phi>_simps(1,5))
finally have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') =
((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (s\<^sub>1 \<star> ?\<gamma>\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
proof -
have "\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') = (s\<^sub>1 \<star> ?\<gamma>\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w]"
using \<gamma>' assoc_naturality [of s\<^sub>1 p\<^sub>0 \<gamma>'] \<open>hseq p\<^sub>0 \<gamma>'\<close> by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = s\<^sub>1 \<star> ?\<gamma>\<^sub>s'"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) \<rho>.T0.antipar(1) comp_assoc_assoc'
by simp
text \<open>Here the fact that \<open>\<phi>\<close> is a retraction is used.\<close>
moreover have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) = cod \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_arr_inv'
whisker_right [of w r\<^sub>0s\<^sub>1.\<phi> "inv r\<^sub>0s\<^sub>1.\<phi>"]
by simp
moreover have "cod (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) = inv r\<^sub>0s\<^sub>1.\<phi> \<star> w"
using \<beta>\<^sub>s_def \<beta>\<^sub>s
by (meson arrI comp_cod_arr seqE)
ultimately show ?thesis
using \<gamma>' \<open>hseq p\<^sub>0 \<gamma>'\<close> comp_arr_dom comp_cod_arr comp_assoc_assoc'
whisker_left [of s\<^sub>1 "p\<^sub>0 \<star> \<gamma>'" "p\<^sub>0 \<star> w"] whisker_left [of p\<^sub>0]
by (metis \<sigma>.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 \<beta>\<^sub>s_def by simp
qed
ultimately have P\<^sub>s': "?P\<^sub>s ?\<gamma>\<^sub>s'"
by simp
have eq\<^sub>s: "\<gamma>\<^sub>s = ?\<gamma>\<^sub>s'"
using 2 \<gamma>\<^sub>s P\<^sub>s' by blast
have "?P\<^sub>t \<gamma>'"
using \<gamma>' comp_cod_arr \<open>\<guillemotleft>p\<^sub>0 \<star> \<gamma>' : p\<^sub>0 \<star> w \<Rightarrow> p\<^sub>0 \<star> w'\<guillemotright>\<close> eq\<^sub>r eq\<^sub>s by auto
thus "\<gamma>' = \<gamma>"
using 3 \<gamma> by blast
qed
qed
ultimately show ?thesis by blast
qed
qed
qed
end
sublocale composite_tabulation_in_maps \<subseteq>
tabulation V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close>
using composite_is_tabulation by simp
sublocale composite_tabulation_in_maps \<subseteq>
tabulation_in_maps V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close>
using T0.is_map \<rho>.leg1_is_map \<rho>.T0.antipar(2) composable \<rho>.leg1_is_map \<rho>.T0.antipar
apply unfold_locales
apply simp
apply (intro left_adjoints_compose)
by auto
subsection "The Classifying Category of Maps"
text \<open>
\sloppypar
We intend to show that if \<open>B\<close> is a bicategory of spans, then \<open>B\<close> is biequivalent to
\<open>Span(Maps(B))\<close>, for a specific category \<open>Maps(B)\<close> derived from \<open>B\<close>.
The category \<open>Maps(B)\<close> is constructed in this section as the ``classifying category'' of
maps of \<open>B\<close>, which has the same objects as \<open>B\<close> and which has as 1-cells the isomorphism classes
of maps of \<open>B\<close>. We show that, if \<open>B\<close> is a bicategory of spans, then \<open>Maps(B)\<close> has pullbacks.
\<close>
locale maps_category =
B: bicategory_of_spans
begin
no_notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>B _\<guillemotright>")
notation B.isomorphic (infix "\<cong>\<^sub>B" 50)
notation B.iso_class ("\<lbrakk>_\<rbrakk>\<^sub>B")
text \<open>
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.
\<close>
text \<open>
The ``hom-categories'' of \<open>Maps(C)\<close> have as arrows the isomorphism classes of maps of \<open>B\<close>.
\<close>
abbreviation Hom
where "Hom a b \<equiv> {F. \<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B}"
lemma in_HomD:
assumes "F \<in> Hom a b"
shows "F \<noteq> {}"
and "B.is_iso_class F"
and "f \<in> F \<Longrightarrow> B.ide f"
and "f \<in> F \<Longrightarrow> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright>"
and "f \<in> F \<Longrightarrow> B.is_left_adjoint f"
and "f \<in> F \<Longrightarrow> F = \<lbrakk>f\<rbrakk>\<^sub>B"
proof -
show "F \<noteq> {}"
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 \<in> F \<Longrightarrow> B.ide f"
using assms 1 B.iso_class_memb_is_ide by blast
obtain f' where f': "\<guillemotleft>f' : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f' \<and> F = \<lbrakk>f'\<rbrakk>\<^sub>B"
using assms by auto
show "f \<in> F \<Longrightarrow> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright>"
using assms f' B.iso_class_def B.isomorphic_implies_hpar by auto
show "f \<in> F \<Longrightarrow> B.is_left_adjoint f"
using assms f' B.iso_class_def B.left_adjoint_preserved_by_iso [of f'] by auto
show "f \<in> F \<Longrightarrow> F = \<lbrakk>f\<rbrakk>\<^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 \<equiv> {h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h)}"
lemma in_CompI [intro]:
assumes "B.is_iso_class F" and "B.is_iso_class G"
and "f \<in> F" and "g \<in> G" and "g \<star> f \<cong>\<^sub>B h"
shows "h \<in> Comp G F"
unfolding Comp_def
using assms by auto
lemma in_CompE [elim]:
assumes "h \<in> Comp G F"
and "\<And>f g. \<lbrakk> B.is_iso_class F; B.is_iso_class G; f \<in> F; g \<in> G; g \<star> f \<cong>\<^sub>B h \<rbrakk> \<Longrightarrow> T"
shows T
using assms Comp_def by auto
lemma is_iso_class_Comp:
assumes "Comp G F \<noteq> {}"
shows "B.is_iso_class (Comp G F)"
proof -
obtain h where h: "h \<in> 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 \<and> B.is_iso_class G \<and> f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using h Comp_def by auto
have "Comp G F = \<lbrakk>g \<star> f\<rbrakk>\<^sub>B \<and> B.ide (g \<star> f)"
proof (intro conjI)
show "B.ide (g \<star> f)"
using fg B.iso_class_memb_is_ide B.isomorphic_implies_ide(1) by auto
show "Comp G F = \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
show "\<lbrakk>g \<star> f\<rbrakk>\<^sub>B \<subseteq> Comp G F"
unfolding Comp_def B.iso_class_def
using fg by auto
show "Comp G F \<subseteq> \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
fix h'
assume h': "h' \<in> Comp G F"
obtain f' g' where f'g': "f' \<in> F \<and> g' \<in> G \<and> g' \<star> f' \<cong>\<^sub>B h'"
using h' Comp_def by auto
have 1: "f' \<cong>\<^sub>B f \<and> g' \<cong>\<^sub>B g"
using f'g' fg B.iso_class_elems_isomorphic by auto
moreover have "B.ide f \<and> B.ide f' \<and> B.ide g \<and> B.ide g'"
using 1 B.isomorphic_implies_hpar by auto
ultimately have "g' \<star> f' \<cong>\<^sub>B g \<star> 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' \<cong>\<^sub>B g \<star> f"
using f'g' B.isomorphic_symmetric B.isomorphic_transitive by blast
thus "h' \<in> B.iso_class (g \<star> 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 \<noteq> {}"
shows "B.is_iso_class F" and "B.is_iso_class G"
and "F \<noteq> {}" and "G \<noteq> {}"
using assms Comp_def by auto
lemma Comp_eqI [intro]:
assumes "h \<in> Comp G F" and "h' \<in> Comp G' F'" and "h \<cong>\<^sub>B h'"
shows "Comp G F = Comp G' F'"
proof -
obtain f g where fg: "f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using assms comp_def by auto
obtain f' g' where f'g': "f' \<in> F' \<and> g' \<in> G' \<and> g' \<star> f' \<cong>\<^sub>B h'"
using assms by auto
have "h \<in> Comp G F \<inter> Comp G' F'"
by (meson IntI assms in_CompE in_CompI B.isomorphic_symmetric B.isomorphic_transitive)
hence "Comp G F \<inter> Comp G' F' \<noteq> {}"
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 \<in> Comp G F"
shows "Comp G F = \<lbrakk>h\<rbrakk>\<^sub>B"
proof
show "Comp G F \<subseteq> \<lbrakk>h\<rbrakk>\<^sub>B"
proof
fix h'
assume h': "h' \<in> Comp G F"
obtain f g where fg: "f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using assms by auto
obtain f' g' where f'g': "f' \<in> F \<and> g' \<in> G \<and> g' \<star> f' \<cong>\<^sub>B h'"
using h' by auto
have "f \<cong>\<^sub>B f' \<and> g \<cong>\<^sub>B g'"
using assms fg f'g' in_HomD(6) B.iso_class_elems_isomorphic by auto
moreover have "B.ide f \<and> B.ide f' \<and> B.ide g \<and> 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 \<and> src g = trg f' \<and> src g' = trg f \<and> 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 \<star> f \<cong>\<^sub>B g' \<star> f'"
using fg f'g' B.hcomp_ide_isomorphic B.hcomp_isomorphic_ide B.isomorphic_transitive
by metis
thus "h' \<in> \<lbrakk>h\<rbrakk>\<^sub>B"
using fg f'g' B.isomorphic_symmetric B.isomorphic_transitive B.iso_class_def [of h]
by blast
qed
show "\<lbrakk>h\<rbrakk>\<^sub>B \<subseteq> Comp G F"
proof (unfold B.iso_class_def Comp_def)
obtain f g where 1: "f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using assms in_HomD Comp_def
by (meson in_CompE)
show "{h'. B.isomorphic h h'} \<subseteq>
{h. B.is_iso_class F \<and> B.is_iso_class G \<and> (\<exists>f g. f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h)}"
using assms 1 B.isomorphic_transitive by blast
qed
qed
interpretation concrete_category \<open>Collect B.obj\<close> Hom B.iso_class \<open>\<lambda>_ _ _. Comp\<close>
proof
show "\<And>a. a \<in> Collect B.obj \<Longrightarrow> \<lbrakk>a\<rbrakk>\<^sub>B \<in> 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 "\<And>a b c F G. \<lbrakk> a \<in> Collect B.obj; b \<in> Collect B.obj; c \<in> Collect B.obj;
F \<in> Hom a b; G \<in> Hom b c \<rbrakk> \<Longrightarrow> Comp G F \<in> Hom a c"
proof -
fix a b c F G
assume a: "a \<in> Collect B.obj" and b: "b \<in> Collect B.obj" and c: "c \<in> Collect B.obj"
and F: "F \<in> Hom a b" and G: "G \<in> Hom b c"
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B"
using F by blast
obtain g where g: "\<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> B.is_left_adjoint g \<and> G = \<lbrakk>g\<rbrakk>\<^sub>B"
using G by blast
have "{h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)} =
\<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
show "{h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)}
\<subseteq> \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
fix h
assume "h \<in> {h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)}"
hence h: "B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)"
by simp
show "h \<in> \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof -
obtain f' g' where f'g': "g' \<in> G \<and> f' \<in> F \<and> g' \<star> f' \<cong>\<^sub>B h"
using h by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : f \<Rightarrow>\<^sub>B f'\<guillemotright> \<and> B.iso \<phi>"
using f f'g' F B.iso_class_def by auto
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : g \<Rightarrow>\<^sub>B g'\<guillemotright> \<and> B.iso \<psi>"
using g f'g' G B.iso_class_def by auto
have 1: "\<guillemotleft>\<psi> \<star> \<phi> : g \<star> f \<Rightarrow>\<^sub>B g' \<star> f'\<guillemotright>"
using f g \<phi> \<psi> by auto
moreover have "B.iso (\<psi> \<star> \<phi>)"
using f g \<phi> \<psi> 1 B.iso_hcomp by auto
ultimately show ?thesis
using f'g' B.iso_class_def B.isomorphic_def by auto
qed
qed
show "\<lbrakk>g \<star> f\<rbrakk>\<^sub>B \<subseteq> {h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^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: "\<And>gf. gf \<in> B.iso_class (g \<star> f) \<Longrightarrow>
B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B gf)"
by blast
show "Comp G F \<in> Hom a c"
proof -
have gf: "B.is_left_adjoint (g \<star> f)"
by (meson f g B.hseqE B.hseqI B.left_adjoints_compose)
obtain gf' where gf': "B.adjoint_pair (g \<star> f) gf'"
using gf by blast
hence "\<lbrakk>g \<star> f\<rbrakk>\<^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 "\<And>a b F. \<lbrakk> a \<in> Collect B.obj; F \<in> Hom a b \<rbrakk> \<Longrightarrow> Comp F \<lbrakk>a\<rbrakk>\<^sub>B = F"
proof -
fix a b F
assume a: "a \<in> Collect B.obj"
assume F: "F \<in> Hom a b"
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B"
using F by auto
have *: "\<And>f'. f' \<in> F \<Longrightarrow> \<guillemotleft>f' : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.ide f' \<and> f \<cong>\<^sub>B f'"
using f B.iso_class_def by force
show "Comp F \<lbrakk>a\<rbrakk>\<^sub>B = F"
proof
show "Comp F \<lbrakk>a\<rbrakk>\<^sub>B \<subseteq> F"
proof
fix h
assume "h \<in> Comp F \<lbrakk>a\<rbrakk>\<^sub>B"
hence h: "\<exists>f' a'. f' \<in> F \<and> a' \<in> \<lbrakk>a\<rbrakk>\<^sub>B \<and> f' \<star> a' \<cong>\<^sub>B h"
unfolding Comp_def by auto
obtain f' a' where f'a': "f' \<in> F \<and> a' \<in> \<lbrakk>a\<rbrakk>\<^sub>B \<and> f' \<star> a' \<cong>\<^sub>B h"
using h by auto
have "B.isomorphic f h"
proof -
have "B.isomorphic f (f \<star> a)"
using f B.iso_runit' [of f] B.isomorphic_def B.left_adjoint_is_ide
by blast
also have "f \<star> a \<cong>\<^sub>B f' \<star> a"
using f f'a' B.iso_class_def B.hcomp_isomorphic_ide
apply (elim conjE B.in_hhomE) by auto
also have "f' \<star> a \<cong>\<^sub>B f' \<star> a'"
using f'a' * [of f'] B.iso_class_def B.hcomp_ide_isomorphic by auto
also have "f' \<star> a' \<cong>\<^sub>B h"
using f'a' by simp
finally show ?thesis by blast
qed
thus "h \<in> F"
using f B.iso_class_def by simp
qed
show "F \<subseteq> Comp F \<lbrakk>a\<rbrakk>\<^sub>B"
proof
fix h
assume h: "h \<in> F"
have "f \<in> F"
using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto
moreover have "a \<in> B.iso_class a"
using a B.iso_class_def B.isomorphic_reflexive by auto
moreover have "f \<star> a \<cong>\<^sub>B h"
proof -
have "f \<star> a \<cong>\<^sub>B f"
using f B.iso_runit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast
also have "f \<cong>\<^sub>B h"
using h * by simp
finally show ?thesis by blast
qed
ultimately show "h \<in> Comp F \<lbrakk>a\<rbrakk>\<^sub>B"
unfolding Comp_def
using a f F B.is_iso_classI B.left_adjoint_is_ide by auto
qed
qed
qed
show "\<And>a b F. \<lbrakk> b \<in> Collect B.obj; F \<in> Hom a b \<rbrakk> \<Longrightarrow> Comp \<lbrakk>b\<rbrakk>\<^sub>B F = F"
proof -
fix a b F
assume b: "b \<in> Collect B.obj"
assume F: "F \<in> Hom a b"
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B"
using F by auto
have *: "\<And>f'. f' \<in> F \<Longrightarrow> \<guillemotleft>f' : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.ide f' \<and> f \<cong>\<^sub>B f'"
using f B.iso_class_def by force
show "Comp \<lbrakk>b\<rbrakk>\<^sub>B F = F"
proof
show "Comp \<lbrakk>b\<rbrakk>\<^sub>B F \<subseteq> F"
proof
fix h
assume "h \<in> Comp \<lbrakk>b\<rbrakk>\<^sub>B F"
hence h: "\<exists>b' f'. b' \<in> \<lbrakk>b\<rbrakk>\<^sub>B \<and> f' \<in> F \<and> b' \<star> f' \<cong>\<^sub>B h"
unfolding Comp_def by auto
obtain b' f' where b'f': "b' \<in> \<lbrakk>b\<rbrakk>\<^sub>B \<and> f' \<in> F \<and> b' \<star> f' \<cong>\<^sub>B h"
using h by auto
have "f \<cong>\<^sub>B h"
proof -
have "f \<cong>\<^sub>B b \<star> f"
using f B.iso_lunit' [of f] B.isomorphic_def B.left_adjoint_is_ide
by blast
also have "... \<cong>\<^sub>B b \<star> f'"
using f b'f' B.iso_class_def B.hcomp_ide_isomorphic
by (elim conjE B.in_hhomE, auto)
also have "... \<cong>\<^sub>B b' \<star> f'"
using b'f' * [of f'] B.iso_class_def B.hcomp_isomorphic_ide by auto
also have "... \<cong>\<^sub>B h"
using b'f' by simp
finally show ?thesis by blast
qed
thus "h \<in> F"
using f B.iso_class_def by simp
qed
show "F \<subseteq> Comp \<lbrakk>b\<rbrakk>\<^sub>B F"
proof
fix h
assume h: "h \<in> F"
have "f \<in> F"
using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto
moreover have "b \<in> B.iso_class b"
using b B.iso_class_def B.isomorphic_reflexive by auto
moreover have "b \<star> f \<cong>\<^sub>B h"
proof -
have "b \<star> f \<cong>\<^sub>B f"
using f B.iso_lunit [of f] B.isomorphic_def B.left_adjoint_is_ide
by blast
also have "f \<cong>\<^sub>B h"
using h * by simp
finally show ?thesis by blast
qed
ultimately show "h \<in> Comp \<lbrakk>b\<rbrakk>\<^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 "\<And>a b c d F G H.
\<lbrakk> a \<in> Collect B.obj; b \<in> Collect B.obj; c \<in> Collect B.obj; d \<in> Collect B.obj;
F \<in> Hom a b; G \<in> Hom b c; H \<in> Hom c d \<rbrakk> \<Longrightarrow>
Comp H (Comp G F) = Comp (Comp H G) F"
proof -
fix a b c d F G H
assume F: "F \<in> Hom a b" and G: "G \<in> Hom b c" and H: "H \<in> Hom c d"
show "Comp H (Comp G F) = Comp (Comp H G) F"
proof
show "Comp H (Comp G F) \<subseteq> Comp (Comp H G) F"
proof
fix x
assume x: "x \<in> Comp H (Comp G F)"
obtain f g h gf
where 1: "B.is_iso_class F \<and> B.is_iso_class G \<and> B.is_iso_class H \<and>
h \<in> H \<and> g \<in> G \<and> f \<in> F \<and> gf \<in> Comp G F \<and> g \<star> f \<cong>\<^sub>B gf \<and> h \<star> gf \<cong>\<^sub>B x"
using x unfolding Comp_def by blast
have hgf: "B.ide f \<and> B.ide g \<and> 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 \<star> g \<star> f \<cong>\<^sub>B x"
by (metis "1" B.hcomp_ide_isomorphic B.hseqE B.ide_char'
B.isomorphic_implies_hpar(4) B.isomorphic_implies_ide(1)
B.isomorphic_transitive hgf)
moreover have "(h \<star> g) \<star> f \<cong>\<^sub>B h \<star> g \<star> 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: "\<guillemotleft>h \<star> g : b \<rightarrow>\<^sub>B d\<guillemotright>"
using G H 1 in_HomD(4) by blast
moreover have "h \<star> g \<in> 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 \<in> Comp (Comp H G) F"
by (metis "1" B.isomorphic_transitive emptyE in_CompI is_iso_class_Comp)
qed
show "Comp (Comp H G) F \<subseteq> Comp H (Comp G F)"
proof
fix x
assume x: "x \<in> Comp (Comp H G) F"
obtain f g h hg
where 1: "B.is_iso_class F \<and> B.is_iso_class G \<and> B.is_iso_class H \<and>
h \<in> H \<and> g \<in> G \<and> f \<in> F \<and> hg \<in> Comp H G \<and> h \<star> g \<cong>\<^sub>B hg \<and> hg \<star> f \<cong>\<^sub>B x"
using x unfolding Comp_def by blast
have hgf: "B.ide f \<and> B.ide g \<and> B.ide h \<and> src h = trg g \<and> 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 \<star> g) \<star> f \<cong>\<^sub>B x"
by (meson "1" B.hcomp_isomorphic_ide B.hseqE B.ideD(1) B.isomorphic_implies_ide(1)
B.isomorphic_symmetric B.isomorphic_transitive hgf)
moreover have "(h \<star> g) \<star> f \<cong>\<^sub>B h \<star> g \<star> f"
using hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def by auto
moreover have "g \<star> f \<in> Comp G F"
using 1 F G hgf B.isomorphic_reflexive by blast
ultimately show "x \<in> Comp H (Comp G F)"
using 1 hgf is_iso_class_Comp [of G F] B.isomorphic_reflexive [of "g \<star> 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 (\<lambda>_ _ _. Comp)"
..
sublocale concrete_category \<open>Collect B.obj\<close> Hom B.iso_class \<open>\<lambda>_ _ _. Comp\<close>
using is_concrete_category by simp
abbreviation comp (infixr "\<odot>" 55)
where "comp \<equiv> COMP"
notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
lemma Map_memb_in_hhom:
assumes "\<guillemotleft>F : A \<rightarrow> B\<guillemotright>" and "f \<in> Map F"
shows "\<guillemotleft>f : Dom A \<rightarrow>\<^sub>B Dom B\<guillemotright>"
proof -
have "\<guillemotleft>f : Dom F \<rightarrow>\<^sub>B Cod F\<guillemotright>"
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 "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright>"
shows "\<guillemotleft>MkArr a b \<lbrakk>f\<rbrakk>\<^sub>B : MkIde a \<rightarrow> MkIde b\<guillemotright>"
using assms MkArr_in_hom by blast
text \<open>
The isomorphisms in \<open>Maps(B)\<close> are the isomorphism classes of equivalence maps in \<open>B\<close>.
\<close>
lemma iso_char:
shows "iso F \<longleftrightarrow> arr F \<and> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
proof
assume F: "iso F"
have "\<And>f. f \<in> Map F \<Longrightarrow> B.equivalence_map f"
proof -
fix f
assume f: "f \<in> Map F"
obtain G where G: "inverse_arrows F G"
using F by auto
obtain g where g: "g \<in> Map G"
using G arr_char [of G] in_HomD(1) by blast
have f: "f \<in> Map F \<and> \<guillemotleft>f : Dom F \<rightarrow>\<^sub>B Cod F\<guillemotright> \<and> B.ide f \<and> 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 \<in> Map G \<and> \<guillemotleft>g : Cod F \<rightarrow>\<^sub>B Dom F\<guillemotright> \<and> B.ide g \<and> 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 \<star> f) \<cong>\<^sub>B g \<star> f"
proof -
have "g \<star> f \<in> Map (G \<odot> F)"
proof -
have 1: "f \<in> Map F \<and> g \<in> Map G \<and> g \<star> f \<cong>\<^sub>B g \<star> f"
using F G f g B.isomorphic_reflexive by force
have 2: "Dom G = Cod F \<and> 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 \<star> f \<in> 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 \<in> Map (G \<odot> 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 \<odot> F) = \<lbrakk>Dom F\<rbrakk>\<^sub>B"
by (simp add: F G comp_inv_arr iso_is_arr)
moreover have "Dom F = src (g \<star> 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 \<star> g \<cong>\<^sub>B trg f"
proof -
have "f \<star> g \<in> Map (F \<odot> G)"
proof -
have 1: "f \<in> Map F \<and> g \<in> Map G \<and> f \<star> g \<cong>\<^sub>B f \<star> g"
using F G f g B.isomorphic_reflexive by force
have 2: "Dom G = Cod F \<and> 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 \<star> g \<in> 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 \<in> Map (F \<odot> 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 \<odot> G) = \<lbrakk>Cod F\<rbrakk>\<^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 \<star> 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 \<and> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
using F by blast
next
assume F: "arr F \<and> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
show "iso F"
proof -
obtain f where f: "f \<in> Map F"
using F arr_char in_HomD(1) by blast
have f_in_hhom: "\<guillemotleft>f : Dom F \<rightarrow>\<^sub>B Cod F\<guillemotright>"
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 \<eta> \<epsilon>' where \<epsilon>': "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'"
using f F B.equivalence_map_def by auto
interpret \<epsilon>': equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using \<epsilon>' by auto
obtain \<epsilon> where \<epsilon>: "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using f F \<epsilon>'.ide_right \<epsilon>'.antipar \<epsilon>'.unit_in_hom \<epsilon>'.unit_is_iso B.equivalence_map_def
B.equivalence_refines_to_adjoint_equivalence [of f g \<eta>]
by auto
interpret \<epsilon>: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using \<epsilon> by simp
have g_in_hhom: "\<guillemotleft>g : Cod F \<rightarrow>\<^sub>B Dom F\<guillemotright>"
using \<epsilon>.ide_right \<epsilon>.antipar f_in_hhom by auto
have fg: "B.quasi_inverses f g"
using B.quasi_inverses_def \<epsilon>.equivalence_in_bicategory_axioms by auto
have g: "\<guillemotleft>g : Cod F \<rightarrow>\<^sub>B Dom F\<guillemotright> \<and> B.is_left_adjoint g \<and> \<lbrakk>g\<rbrakk>\<^sub>B = \<lbrakk>g\<rbrakk>\<^sub>B"
using \<epsilon>'.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) \<lbrakk>f\<rbrakk>\<^sub>B"
using F MkArr_Map \<open>Map F = B.iso_class f\<close> 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) \<lbrakk>g\<rbrakk>\<^sub>B"
have "arr ?G"
using MkArr_in_hom' g by blast
hence G_in_hom: "\<guillemotleft>?G : MkIde (Cod F) \<rightarrow> MkIde (Dom F)\<guillemotright>"
using arr_char MkArr_in_hom by simp
have "inverse_arrows F ?G"
proof
show "ide (?G \<odot> F)"
proof -
have "?G \<odot> 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 \<odot> F) = Dom (dom F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Cod (?G \<odot> F) = Cod (dom F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Map (?G \<odot> F) = Map (dom F)"
proof -
have "Map (?G \<odot> F) = Comp \<lbrakk>g\<rbrakk>\<^sub>B \<lbrakk>f\<rbrakk>\<^sub>B"
using 1 comp_char [of ?G F] \<open>Map F = B.iso_class f\<close> by simp
also have "... = \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof -
have "g \<star> f \<in> Comp \<lbrakk>g\<rbrakk>\<^sub>B \<lbrakk>f\<rbrakk>\<^sub>B"
by (metis \<epsilon>.ide_left \<epsilon>.ide_right \<epsilon>.unit_in_vhom \<epsilon>.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
\<open>Map F = B.iso_class f\<close>
by auto
qed
also have "... = \<lbrakk>src f\<rbrakk>\<^sub>B"
using \<epsilon>.unit_in_hom \<epsilon>.unit_is_iso B.isomorphic_def B.iso_class_def
B.isomorphic_symmetric
by (intro B.iso_class_eqI, blast)
also have "... = \<lbrakk>Dom F\<rbrakk>\<^sub>B"
using \<epsilon>.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 \<odot> ?G)"
proof -
have "F \<odot> ?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 \<odot> ?G) = Dom (cod F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Cod (F \<odot> ?G) = Cod (cod F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Map (F \<odot> ?G) = Map (cod F)"
proof -
have "Map (F \<odot> ?G) = Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
using 1 comp_char [of F ?G] \<open>Map F = \<lbrakk>f\<rbrakk>\<^sub>B\<close> by simp
also have "... = \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
proof -
have "f \<star> g \<in> Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
by (metis \<epsilon>.antipar(1) \<epsilon>.ide_left \<epsilon>.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
\<open>Map F = \<lbrakk>f\<rbrakk>\<^sub>B\<close>
by auto
qed
also have "... = \<lbrakk>trg f\<rbrakk>\<^sub>B"
proof -
have "trg f \<in> \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
using \<epsilon>.counit_in_hom \<epsilon>.counit_is_iso B.isomorphic_def B.iso_class_def
B.isomorphic_symmetric
by blast
thus ?thesis
using B.iso_class_eqI
by (metis \<epsilon>.antipar(1) \<epsilon>.ide_left \<epsilon>.ide_right B.ide_hcomp
B.ide_in_iso_class B.is_iso_classI B.iso_class_elems_isomorphic)
qed
also have "... = \<lbrakk>Cod F\<rbrakk>\<^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 \<longleftrightarrow> arr F \<and> (\<exists>f. f \<in> Map F \<and> B.equivalence_map f)"
proof -
have "arr F \<Longrightarrow> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f) \<longleftrightarrow>
(\<exists>f. f \<in> Map F \<and> B.equivalence_map f)"
proof
assume F: "arr F"
show "(\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f) \<Longrightarrow>
(\<exists>f. f \<in> Map F \<and> B.equivalence_map f)"
using F arr_char in_HomD(1) by blast
show "(\<exists>f. f \<in> Map F \<and> B.equivalence_map f) \<Longrightarrow>
(\<forall>f. f \<in> Map F \<longrightarrow> 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 \<open>
The following mapping takes a map in \<open>B\<close> to the corresponding arrow of \<open>Maps(B)\<close>.
\<close>
abbreviation CLS ("\<lbrakk>\<lbrakk>_\<rbrakk>\<rbrakk>")
where "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<equiv> MkArr (src f) (trg f) \<lbrakk>f\<rbrakk>\<^sub>B"
lemma ide_CLS_obj:
assumes "B.obj a"
shows "ide \<lbrakk>\<lbrakk>a\<rbrakk>\<rbrakk>"
by (simp add: assms B.obj_simps)
lemma CLS_in_hom:
assumes "B.is_left_adjoint f"
shows "\<guillemotleft>\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> : \<lbrakk>\<lbrakk>src f\<rbrakk>\<rbrakk> \<rightarrow> \<lbrakk>\<lbrakk>trg f\<rbrakk>\<rbrakk>\<guillemotright>"
using assms B.left_adjoint_is_ide MkArr_in_hom MkArr_in_hom' by simp
lemma CLS_eqI:
assumes "B.ide f"
shows "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk> \<longleftrightarrow> f \<cong>\<^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 "\<lbrakk>\<lbrakk>f \<star> g\<rbrakk>\<rbrakk> = MkArr (src g) (trg f) (Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B)"
proof -
have "\<lbrakk>\<lbrakk>f \<star> g\<rbrakk>\<rbrakk> = MkArr (src g) (trg f) \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
using assms B.left_adjoint_is_ide by simp
moreover have "\<lbrakk>f \<star> g\<rbrakk>\<^sub>B = Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
proof
show "\<lbrakk>f \<star> g\<rbrakk>\<^sub>B \<subseteq> Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^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 \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B \<subseteq> \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
by (metis Comp_eq_iso_class_memb \<open>\<lbrakk>f \<star> g\<rbrakk>\<^sub>B \<subseteq> Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B\<close>
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 \<star> g \<cong>\<^sub>B h"
shows "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>h\<rbrakk>\<rbrakk>"
proof -
have hseq_fg: "B.hseq f g"
using assms B.isomorphic_implies_hpar(1) by simp
have "seq \<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk>"
using assms hseq_fg CLS_in_hom [of f] CLS_in_hom [of g]
apply (elim B.hseqE) by auto
hence "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk> = MkArr (src g) (trg f) (Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B)"
using comp_char [of "CLS f" "CLS g"] by simp
also have "... = \<lbrakk>\<lbrakk>f \<star> g\<rbrakk>\<rbrakk>"
using assms hseq_fg CLS_hcomp
apply (elim B.hseqE)
using B.adjoint_pair_antipar(1) by auto
also have "... = \<lbrakk>\<lbrakk>h\<rbrakk>\<rbrakk>"
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 \<open>
The following mapping that takes an arrow of \<open>Maps(B)\<close> and chooses some
representative map in \<open>B\<close>.
\<close>
definition REP
where "REP F \<equiv> if arr F then SOME f. f \<in> Map F else B.null"
lemma REP_in_Map:
assumes "arr A"
shows "REP A \<in> Map A"
proof -
have "Map A \<noteq> {}"
using assms arr_char in_HomD(1) by blast
thus ?thesis
using assms REP_def someI_ex [of "\<lambda>f. f \<in> Map A"] by auto
qed
lemma REP_in_hhom [intro]:
assumes "in_hom F A B"
shows "\<guillemotleft>REP F : src (REP A) \<rightarrow>\<^sub>B src (REP B)\<guillemotright>"
and "B.is_left_adjoint (REP F)"
proof -
have "Map F \<noteq> {}"
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 \<in> Map F"
using assms REP_def someI_ex [of "\<lambda>f. f \<in> 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 "\<guillemotleft>REP F : Dom A \<rightarrow>\<^sub>B Dom B\<guillemotright>"
using assms 1 in_hom_char [of F A B] Map_memb_in_hhom by auto
thus "\<guillemotleft>REP F : src (REP A) \<rightarrow>\<^sub>B src (REP B)\<guillemotright>"
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 \<in> \<lbrakk>REP F\<rbrakk>\<^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 \<cong>\<^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 \<cong>\<^sub>B trg (REP A)"
using assms ide_char isomorphic_REP_src by auto
lemma CLS_REP:
assumes "arr F"
shows "\<lbrakk>\<lbrakk>REP F\<rbrakk>\<rbrakk> = 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 \<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<cong>\<^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 \<star> REP G \<cong>\<^sub>B REP (F \<odot> 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 \<odot> G) = Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B"
using assms seq_char null_char
by (metis (no_types, lifting) CLS_REP Map.simps(1) Map_comp)
have "Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B = \<lbrakk>REP F \<star> REP G\<rbrakk>\<^sub>B"
proof -
have "REP F \<star> REP G \<in> Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B"
proof -
have "REP F \<in> Map F \<and> REP G \<in> Map G"
using assms seq_char REP_in_Map by auto
moreover have "REP F \<star> REP G \<cong>\<^sub>B REP F \<star> 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 "\<lbrakk>REP F\<rbrakk>\<^sub>B \<in> Hom (Cod G) (Cod F)"
using assms 1 2 arr_char [of F] by auto
moreover have "\<lbrakk>REP G\<rbrakk>\<^sub>B \<in> 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 \<odot> G) \<in> Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B"
proof -
have "Map (F \<odot> 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 \<odot> G"] by simp
qed
ultimately show ?thesis
using B.iso_class_def by simp
qed
text \<open>
We now show that \<open>Maps(B)\<close> has pullbacks. For this we need to exhibit
functions \<open>PRJ\<^sub>0\<close> and \<open>PRJ\<^sub>1\<close> that produce the legs of the pullback of a cospan \<open>(H, K)\<close>
and verify that they have the required universal property. These are obtained by
choosing representatives \<open>h\<close> and \<open>k\<close> of \<open>H\<close> and \<open>K\<close>, respectively, and then taking
\<open>PRJ\<^sub>0 = CLS (tab\<^sub>0 (k\<^sup>* \<star> h))\<close> and \<open>PRJ\<^sub>1 = CLS (tab\<^sub>1 (k\<^sup>* \<star> h))\<close>. That these constitute a
pullback in \<open>Maps(B)\<close> follows from the fact that \<open>tab\<^sub>0 (k\<^sup>* \<star> h)\<close> and \<open>tab\<^sub>1 (k\<^sup>* \<star> h)\<close>
form a pseudo-pullback of \<open>(h, k)\<close> in the underlying bicategory.
For our purposes here, it is not sufficient simply to show that \<open>Maps(B)\<close> 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 \<open>k\<^sup>* \<star> h\<close> 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 \<open>B\<close> to \<open>Span(Maps(B))\<close>, because the components will go from the
apex of a composite span (obtained by pullback) to the apex of a tabulation (chosen
arbitrarily using \<open>BS2\<close>) and these need not be in agreement with each other.
\<close>
definition PRJ\<^sub>0
where "PRJ\<^sub>0 \<equiv> \<lambda>K H. if cospan K H then \<lbrakk>\<lbrakk>B.tab\<^sub>0 ((REP K)\<^sup>* \<star> (REP H))\<rbrakk>\<rbrakk> else null"
definition PRJ\<^sub>1
where "PRJ\<^sub>1 \<equiv> \<lambda>K H. if cospan K H then \<lbrakk>\<lbrakk>B.tab\<^sub>1 ((REP K)\<^sup>* \<star> (REP H))\<rbrakk>\<rbrakk> else null"
interpretation elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1
proof
show "\<And>H K. \<not> cospan K H \<Longrightarrow> PRJ\<^sub>0 K H = null"
unfolding PRJ\<^sub>0_def by auto
show "\<And>H K. \<not> cospan K H \<Longrightarrow> PRJ\<^sub>1 K H = null"
unfolding PRJ\<^sub>1_def by auto
show "\<And>H K. cospan K H \<Longrightarrow> 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 \<in> Map H"
using h_def HK REP_in_Map by blast
have k: "k \<in> Map K"
using k_def HK REP_in_Map by blast
have 1: "B.is_left_adjoint h \<and> B.is_left_adjoint k \<and> B.ide h \<and> B.ide k \<and> 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 \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h
using 1 by unfold_locales auto
interpret k: map_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg k
using 1 by unfold_locales auto
interpret hk: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h k
using 1 by unfold_locales auto
let ?f = "B.tab\<^sub>0 (k\<^sup>* \<star> h)"
let ?g = "B.tab\<^sub>1 (k\<^sup>* \<star> h)"
have span: "span \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
using dom_char CLS_in_hom [of ?f] CLS_in_hom [of ?g] by auto
have seq: "seq H \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
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 \<odot> PRJ\<^sub>1 K H = H \<odot> PRJ\<^sub>0 K H"
proof -
have iso: "h \<star> ?f \<cong>\<^sub>B k \<star> ?g"
using hk.\<phi>_uniqueness B.isomorphic_symmetric B.isomorphic_def by blast
have *: "Comp (Map H) \<lbrakk>?f\<rbrakk>\<^sub>B = Comp (Map K) \<lbrakk>?g\<rbrakk>\<^sub>B"
proof (intro Comp_eqI)
show "h \<star> ?f \<in> Comp (Map H) \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
proof (unfold Comp_def)
have "B.is_iso_class \<lbrakk>?f\<rbrakk>\<^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 \<in> \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
by (simp add: B.ide_in_iso_class(1))
moreover have "\<guillemotleft>?f : src (B.tab\<^sub>0 (k\<^sup>* \<star> h)) \<rightarrow>\<^sub>B Dom H\<guillemotright>"
using seq seq_char by simp
moreover have "h \<in> Map H"
by fact
moreover have "\<guillemotleft>h : Dom H \<rightarrow>\<^sub>B Cod H\<guillemotright>"
by (simp add: HK h_def)
moreover have "h \<star> ?f \<cong>\<^sub>B h \<star> ?f"
using B.isomorphic_reflexive by auto
ultimately show "h \<star> B.tab\<^sub>0 (k\<^sup>* \<star> h)
\<in> {h'. B.is_iso_class \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
B.is_iso_class (Map H) \<and>
(\<exists>f g. f \<in> \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
g \<in> Map H \<and> g \<star> f \<cong>\<^sub>B h')}"
by auto
qed
show "k \<star> ?g \<in> Comp (Map K) \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
proof (unfold Comp_def)
have "B.is_iso_class \<lbrakk>?g\<rbrakk>\<^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 \<in> \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
by (simp add: B.ide_in_iso_class(1))
moreover have "\<guillemotleft>?g : src (B.tab\<^sub>1 (k\<^sup>* \<star> h)) \<rightarrow>\<^sub>B Dom K\<guillemotright>"
using seq seq_char B.in_hhom_def seq' by auto
moreover have "k \<in> Map K"
by fact
moreover have "\<guillemotleft>k : Dom K \<rightarrow>\<^sub>B Cod K\<guillemotright>"
by (simp add: HK k_def)
moreover have "k \<star> ?g \<cong>\<^sub>B k \<star> ?g"
using B.isomorphic_reflexive iso B.isomorphic_implies_hpar(2) by auto
ultimately show "k \<star> B.tab\<^sub>1 (k\<^sup>* \<star> h)
\<in> {h'. B.is_iso_class \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
B.is_iso_class (Map K) \<and>
(\<exists>f g. f \<in> \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
g \<in> Map K \<and> g \<star> f \<cong>\<^sub>B h')}"
by auto
qed
show "h \<star> ?f \<cong>\<^sub>B k \<star> ?g"
using iso by simp
qed
have "K \<odot> PRJ\<^sub>1 K H = K \<odot> \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
unfolding PRJ\<^sub>1_def using HK h_def k_def by simp
also have "... = MkArr (src ?g) (Cod K) (Comp (Map K) \<lbrakk>?g\<rbrakk>\<^sub>B)"
using seq' comp_char [of K "\<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"] by simp
also have "... = MkArr (src ?f) (Cod H) (Comp (Map H) \<lbrakk>?f\<rbrakk>\<^sub>B)"
using * HK cod_char by auto
also have "... = comp H \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"
using seq comp_char [of H "\<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"] 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 "\<And>H K U V. commutative_square K H V U \<Longrightarrow>
\<exists>!E. comp (PRJ\<^sub>1 K H) E = V \<and> 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 \<in> Map H"
using h_def HK REP_in_Map by blast
have k: "k \<in> Map K"
using k_def HK REP_in_Map by blast
have 1: "B.is_left_adjoint h \<and> B.is_left_adjoint k \<and> B.ide h \<and> B.ide k \<and> 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 \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h
using 1 by unfold_locales auto
interpret k: map_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg k
using 1 by unfold_locales auto
interpret hk: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h k
using 1 by unfold_locales auto
let ?f = "B.tab\<^sub>0 (k\<^sup>* \<star> h)"
let ?g = "B.tab\<^sub>1 (k\<^sup>* \<star> 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 \<star> ?u \<cong>\<^sub>B k \<star> ?v"
proof -
have "h \<star> ?u \<cong>\<^sub>B REP (H \<odot> U)"
proof -
have "h \<star> ?u \<cong>\<^sub>B REP H \<star> ?u"
proof -
have "h \<cong>\<^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 "... \<cong>\<^sub>B REP (H \<odot> U)"
using seq_HU isomorphic_hcomp_REP isomorphic_def by blast
finally show ?thesis by blast
qed
also have "... \<cong>\<^sub>B REP (K \<odot> V)"
using seq_HU cs B.isomorphic_reflexive by auto
also have "... \<cong>\<^sub>B (k \<star> ?v)"
proof -
have "... \<cong>\<^sub>B REP K \<star> ?v"
using seq_KV isomorphic_hcomp_REP B.isomorphic_def B.isomorphic_symmetric
by blast
also have "... \<cong>\<^sub>B k \<star> ?v"
proof -
have "k \<cong>\<^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 \<and> ?f \<star> w \<cong>\<^sub>B ?u \<and> ?v \<cong>\<^sub>B (?g \<star> w)"
using * u_is_map v_is_map hk.has_pseudo_pullback [of ?u ?v] B.isomorphic_symmetric
by blast
have w_in_hom: "\<guillemotleft>w : src ?u \<rightarrow>\<^sub>B src ?f\<guillemotright> \<and> 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: "\<guillemotleft>?W : dom U \<rightarrow> dom \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>\<guillemotright>"
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 \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"
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 "\<exists>!E. PRJ\<^sub>1 K H \<odot> E = V \<and> PRJ\<^sub>0 K H \<odot> E = U"
proof -
have "PRJ\<^sub>1 K H \<odot> ?W = V \<and> PRJ\<^sub>0 K H \<odot> ?W = U"
proof -
have "\<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<odot> ?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 "\<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk> \<odot> ?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
"\<And>W'. PRJ\<^sub>1 K H \<odot> W' = V \<and> PRJ\<^sub>0 K H \<odot> W' = U \<Longrightarrow> W' = ?W"
proof -
fix W'
assume "PRJ\<^sub>1 K H \<odot> W' = V \<and> PRJ\<^sub>0 K H \<odot> W' = U"
hence W': "\<guillemotleft>W' : dom U \<rightarrow> dom \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>\<guillemotright> \<and> \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<odot> W' = U \<and> \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk> \<odot> 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 \<star> ?w' \<cong>\<^sub>B ?u"
proof -
have "?f \<star> ?w' \<cong>\<^sub>B REP \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<star> ?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 \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<star> ?w' \<cong>\<^sub>B ?u"
using W' isomorphic_hcomp_REP cs by blast
finally show ?thesis by blast
qed
have gw'_iso_v: "?g \<star> ?w' \<cong>\<^sub>B ?v"
proof -
have "?g \<star> ?w' \<cong>\<^sub>B REP \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk> \<star> ?w'"
proof -
have "?g \<cong>\<^sub>B REP \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
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 "... \<cong>\<^sub>B ?v"
using W' isomorphic_hcomp_REP cs by blast
finally show ?thesis by blast
qed
show "W' = ?W"
proof -
have "W' = \<lbrakk>\<lbrakk>?w'\<rbrakk>\<rbrakk>"
using w' W' CLS_REP by auto
also have "... = ?W"
proof -
have "?w' \<cong>\<^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 \<open>
Here we relate the projections of the chosen pullbacks in \<open>Maps(B)\<close> to the
projections associated with the chosen tabulations in \<open>B\<close>.
\<close>
context composite_tabulation_in_maps
begin
interpretation Maps: maps_category V H \<a> \<i> src trg
..
interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg s\<^sub>1 r\<^sub>0
using \<rho>.leg0_is_map \<sigma>.leg1_is_map composable by unfold_locales auto
lemma prj_char:
shows "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>0 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
and "Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>1 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "Maps.arr (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
using \<sigma>.leg1_in_hom Maps.CLS_in_hom \<sigma>.leg1_is_map Maps.arr_char by auto
moreover have "Maps.arr (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>)"
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) \<lbrakk>r\<^sub>0\<rbrakk>) =
Maps.cod (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
unfolding Maps.arr_char
using \<sigma>.leg1_in_hom \<rho>.leg0_in_hom
by (simp add: Maps.cod_char calculation(1) calculation(2))
ultimately have "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>))\<rbrakk>\<rbrakk> \<and>
Maps.PRJ\<^sub>1 (Maps.CLS r\<^sub>0) (Maps.CLS s\<^sub>1) =
\<lbrakk>\<lbrakk>tab\<^sub>1 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>))\<rbrakk>\<rbrakk>"
unfolding Maps.PRJ\<^sub>0_def Maps.PRJ\<^sub>1_def
using Maps.CLS_in_hom \<sigma>.leg1_is_map \<rho>.leg0_is_map composable by simp
moreover have "r\<^sub>0\<^sup>* \<star> s\<^sub>1 \<cong> (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
proof -
have "r\<^sub>0 \<cong> Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>)"
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) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>*"
using \<rho>.leg0_is_map
by (simp add: isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint)
moreover have 4: "s\<^sub>1 \<cong> Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
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 \<rho>.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) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* = trg s\<^sub>1"
by (metis 3 \<rho>.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 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>0 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk> \<and>
Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>1 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
using r\<^sub>0s\<^sub>1.isomorphic_implies_same_tab by simp
show "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>0 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
using 1 by simp
show "Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>1 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
using 1 by simp
qed
end
context identity_in_bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
notation isomorphic (infix "\<cong>" 50)
text \<open>
A 1-cell \<open>r\<close> in a bicategory of spans is a map if and only if the ``input leg''
\<open>tab\<^sub>0 r\<close> of the chosen tabulation of \<open>r\<close> is an equivalence map.
Since a tabulation of \<open>r\<close> is unique up to equivalence, and equivalence maps compose,
the result actually holds if ``chosen tabulation'' is replaced by ``any tabulation''.
\<close>
lemma is_map_iff_tab\<^sub>0_is_equivalence:
shows "is_left_adjoint r \<longleftrightarrow> 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' \<eta>' \<epsilon>' where \<eta>'\<epsilon>': "equivalence_in_bicategory V H \<a> \<i> src trg (tab\<^sub>0 r) g' \<eta>' \<epsilon>'"
using 1 equivalence_map_def by auto
have "adjoint_pair (tab\<^sub>0 r) g'"
using \<eta>'\<epsilon>' 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' \<cong> (tab\<^sub>0 r)\<^sup>*"
using left_adjoint_determines_right_up_to_iso by simp
thus ?thesis
using \<eta>'\<epsilon>' quasi_inverses_def quasi_inverses_isomorphic_right by blast
qed
obtain \<eta>' \<epsilon>' where \<eta>'\<epsilon>': "equivalence_in_bicategory V H \<a> \<i> src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \<eta>' \<epsilon>'"
using 2 quasi_inverses_def by auto
interpret \<eta>'\<epsilon>': equivalence_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>0 r\<close> \<open>(tab\<^sub>0 r)\<^sup>*\<close> \<eta>' \<epsilon>'
using \<eta>'\<epsilon>' 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 \<star> (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 \<star> (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>* \<star> tab\<^sub>0 r) \<and> is_left_adjoint (tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*)"
using left_adjoints_compose T0.antipar by simp
hence 3: "iso \<eta> \<and> iso \<epsilon>"
using BS3 [of "src (tab\<^sub>0 r)" "(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>0 r" \<eta> \<eta>]
BS3 [of "tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*" "trg (tab\<^sub>0 r)" \<epsilon> \<epsilon>]
T0.unit_in_hom T0.counit_in_hom obj_is_self_adjoint
by auto
hence "equivalence_in_bicategory V H \<a> \<i> src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \<eta> \<epsilon>"
apply unfold_locales by auto
thus "equivalence_map (tab\<^sub>0 r)"
using equivalence_map_def by blast
qed
text \<open>
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.
\<close>
lemma obj_has_symmetric_tab:
assumes "obj r"
shows "tab\<^sub>0 r \<cong> tab\<^sub>1 r"
proof -
have "tab\<^sub>0 r \<cong> r \<star> tab\<^sub>0 r"
proof -
have "trg (tab\<^sub>0 r) = r"
using assms by auto
moreover have "\<guillemotleft>\<l>\<^sup>-\<^sup>1[tab\<^sub>0 r] : tab\<^sub>0 r \<Rightarrow> trg (tab\<^sub>0 r) \<star> tab\<^sub>0 r\<guillemotright> \<and> iso \<l>\<^sup>-\<^sup>1[tab\<^sub>0 r]"
using assms by simp
ultimately show ?thesis
unfolding isomorphic_def by metis
qed
also have "... \<cong> tab\<^sub>1 r"
proof -
have "\<guillemotleft>tab : tab\<^sub>1 r \<Rightarrow> r \<star> tab\<^sub>0 r\<guillemotright>"
using tab_in_hom by simp
moreover have "is_left_adjoint (r \<star> 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 \<star> tab\<^sub>0 r" tab tab] isomorphic_symmetric isomorphic_def
by auto
qed
finally show ?thesis by simp
qed
text \<open>
The chosen tabulation of \<open>r\<close> determines a span in \<open>Maps(B)\<close>.
\<close>
lemma determines_span:
assumes "ide r"
shows "span_in_category Maps.comp \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>\<rparr>"
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 \<open>
Here we consider the situation of two tabulations: a tabulation \<open>\<rho>\<close> of \<open>r\<close>
and a tabulation \<open>\<sigma>\<close> of \<open>s\<close>, both ``legs'' of each tabulation being maps,
together with an arbitrary 2-cell \<open>\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>\<close>.
The 2-cell \<open>\<mu>\<close> at the base composes with the tabulation \<open>\<rho>\<close> to yield a 2-cell
\<open>\<Delta> = (\<mu> \<star> r\<^sub>0) \<cdot> \<rho>\<close> ``over'' s. By property \<open>T1\<close> of tabulation \<open>\<sigma>\<close>, this induces a map
from the apex of \<open>\<rho>\<close> to the apex of \<open>\<sigma>\<close>, which together with the other data
forms a triangular prism whose sides commute up to (unique) isomorphism.
\<close>
text \<open>
$$
\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]
}
$$
\<close>
locale arrow_of_tabulations_in_maps =
bicategory_of_spans V H \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and \<mu> :: 'a +
assumes in_hom: "\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>"
begin
abbreviation (input) \<Delta>
where "\<Delta> \<equiv> (\<mu> \<star> r\<^sub>0) \<cdot> \<rho>"
lemma \<Delta>_in_hom [intro]:
shows "\<guillemotleft>\<Delta> : src \<rho> \<rightarrow> trg \<sigma>\<guillemotright>"
and "\<guillemotleft>\<Delta> : r\<^sub>1 \<Rightarrow> s \<star> r\<^sub>0\<guillemotright>"
proof -
show "\<guillemotleft>\<Delta> : r\<^sub>1 \<Rightarrow> s \<star> r\<^sub>0\<guillemotright>"
using in_hom \<rho>.leg0_in_hom(2) \<rho>.tab_in_vhom' by auto
thus "\<guillemotleft>\<Delta> : src \<rho> \<rightarrow> trg \<sigma>\<guillemotright>"
by (metis \<rho>.tab_simps(3) \<rho>.base_in_hom(2) \<sigma>.tab_simps(3) \<sigma>.base_in_hom(2) arrI in_hom
seqI' vcomp_in_hhom vseq_implies_hpar(1-2))
qed
lemma \<Delta>_simps [simp]:
shows "arr \<Delta>"
and "src \<Delta> = src \<rho>" and "trg \<Delta> = trg \<sigma>"
and "dom \<Delta> = r\<^sub>1" and "cod \<Delta> = s \<star> r\<^sub>0"
using \<Delta>_in_hom by auto
abbreviation is_induced_map
where "is_induced_map w \<equiv> \<sigma>.is_induced_by_cell w r\<^sub>0 \<Delta>"
text \<open>
The following is an equivalent restatement, in elementary terms, of the conditions
for being an induced map.
\<close>
abbreviation (input) is_induced_map'
where "is_induced_map' w \<equiv>
ide w \<and>
(\<exists>\<nu> \<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<Delta> = (s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu>)"
lemma is_induced_map_iff:
shows "is_induced_map w \<longleftrightarrow> is_induced_map' w"
proof
assume w: "is_induced_map' w"
show "is_induced_map w"
proof
have 1: "dom \<Delta> = r\<^sub>1"
by auto
interpret w: arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg
r\<^sub>0 \<open>dom \<Delta>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w
proof -
have "arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 w"
using w apply unfold_locales by auto
thus "arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) s \<sigma> s\<^sub>0 s\<^sub>1 w"
using 1 by simp
qed
show "arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) s\<^sub>0 s\<^sub>1 w"
using w.arrow_of_spans_of_maps_axioms by auto
show "\<sigma>.composite_cell w w.the_\<theta> \<cdot> w.the_\<nu> = \<Delta>"
proof -
obtain \<theta> \<nu>
where \<theta>\<nu>: "\<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<Delta> = (s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu>"
using w w.the_\<theta>_props(1) by auto
have "(s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu> = \<Delta>"
using \<theta>\<nu> by argo
moreover have "\<theta> = w.the_\<theta> \<and> \<nu> = w.the_\<nu>"
using \<theta>\<nu> 1 w.the_\<nu>_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 \<a> \<i> 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 \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 w
..
have "dom \<Delta> = r\<^sub>1" by auto
thus ?thesis
using w comp_assoc w.the_\<nu>_props(1) w.the_\<nu>_props(2) w.uw\<theta> by metis
qed
qed
lemma exists_induced_map:
shows "\<exists>w. is_induced_map w"
proof -
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<Delta> = (s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu>"
using \<Delta>_in_hom \<rho>.ide_leg0 \<sigma>.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 \<cong> w'"
using assms \<sigma>.induced_map_unique by blast
definition chine
where "chine \<equiv> 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 "\<guillemotleft>chine : src r\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright>"
and "\<guillemotleft>chine: chine \<Rightarrow> chine\<guillemotright>"
proof -
show "\<guillemotleft>chine : src r\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright>"
using chine_is_induced_map
by (metis \<Delta>_simps(1) \<Delta>_simps(4) \<rho>.leg1_simps(3) \<sigma>.ide_base \<sigma>.ide_leg0 \<sigma>.leg0_simps(3)
\<sigma>.tab_simps(2) arrow_of_spans_of_maps.is_ide arrow_of_spans_of_maps.the_\<nu>_simps(2)
assoc_simps(2) hseqE in_hhom_def seqE src_vcomp vseq_implies_hpar(1))
show "\<guillemotleft>chine: chine \<Rightarrow> chine\<guillemotright>"
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 \<subseteq>
arrow_of_spans_of_maps V H \<a> \<i> 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 \<subseteq>
arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 chine
..
context arrow_of_tabulations_in_maps
begin
text \<open>
The two factorizations of the composite 2-cell \<open>\<Delta>\<close> amount to a naturality condition.
\<close>
lemma \<Delta>_naturality:
shows "(\<mu> \<star> r\<^sub>0) \<cdot> \<rho> = (s \<star> the_\<theta>) \<cdot> \<a>[s, s\<^sub>0, chine] \<cdot> (\<sigma> \<star> chine) \<cdot> the_\<nu>"
using chine_is_induced_map is_induced_map_iff
by (metis leg0_uniquely_isomorphic(2) leg1_uniquely_isomorphic(2) the_\<nu>_props(1) uw\<theta>)
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 \<a> \<i> 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 \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 w
..
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow> w'\<guillemotright> \<and> iso \<phi>"
using assms(2) isomorphic_def by auto
show ?thesis
proof
interpret w': arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<Delta>\<close> s\<^sub>0 s\<^sub>1 w'
proof
show "is_left_adjoint r\<^sub>0"
by (simp add: \<rho>.satisfies_T0)
show "is_left_adjoint (dom \<Delta>)"
by (simp add: \<rho>.leg1_is_map)
show "ide w'" using assms by force
show "\<exists>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright>"
using \<phi> w.the_\<theta>_props \<sigma>.leg0_in_hom(2) assms(2) comp_in_homI hcomp_in_vhom
inv_in_hom isomorphic_implies_hpar(4) w.the_\<theta>_simps(4) w.w_simps(4)
by metis
have "\<guillemotleft>(s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso ((s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>)"
proof (intro conjI)
show "\<guillemotleft>(s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright>"
using \<phi> w.the_\<nu>_props
by (intro comp_in_homI, auto)
thus "iso ((s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>)"
using \<phi> w.the_\<nu>_props
by (meson \<sigma>.ide_leg1 arrI iso_hcomp hseqE ide_is_iso isos_compose seqE)
qed
hence "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>"
by auto
thus "\<exists>\<nu>. \<guillemotleft>\<nu> : dom \<Delta> \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>"
by auto
qed
interpret w': arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg
r\<^sub>0 \<open>dom \<Delta>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w'
..
show "arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) s\<^sub>0 s\<^sub>1 w'"
using w'.arrow_of_spans_of_maps_axioms by auto
show "\<sigma>.composite_cell w' w'.the_\<theta> \<cdot> w'.the_\<nu> = \<Delta>"
proof -
have 1: "w'.the_\<theta> = w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>)"
proof -
have "\<guillemotleft>w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>) : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright>"
using w.the_\<theta>_props \<phi>
by (intro comp_in_homI, auto)
moreover have "\<guillemotleft>w'.the_\<theta> : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright>"
using w'.the_\<theta>_props by simp
ultimately show ?thesis
using w'.leg0_uniquely_isomorphic(2) by blast
qed
moreover have "w'.the_\<nu> = (s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>"
proof -
have "\<guillemotleft>(s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu> : dom \<Delta> \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright>"
using w.the_\<nu>_props \<phi>
by (intro comp_in_homI, auto)
moreover have "iso ((s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>)"
using w.the_\<nu>_props \<phi> iso_hcomp
by (meson \<sigma>.ide_leg1 arrI calculation hseqE ide_is_iso isos_compose seqE)
ultimately show ?thesis
using w'.the_\<nu>_props w'.leg1_uniquely_isomorphic(2) by blast
qed
ultimately have "\<sigma>.composite_cell w' w'.the_\<theta> \<cdot> w'.the_\<nu> =
\<sigma>.composite_cell w' (w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>)) \<cdot> (s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>"
by simp
also have "... = (s \<star> w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>)) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot>
(\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>"
using comp_assoc by presburger
also have "... = (s \<star> w.the_\<theta>) \<cdot> ((s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot>
(\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>)) \<cdot> w.the_\<nu>"
using 1 comp_assoc w'.the_\<theta>_simps(1) whisker_left
by auto
also have "... = (s \<star> w.the_\<theta>) \<cdot> (\<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w)) \<cdot> w.the_\<nu>"
proof -
have "(s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot> (\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>) =
\<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w)"
proof -
have "(s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot> (\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>) =
\<a>[s, s\<^sub>0, w] \<cdot> ((s \<star> s\<^sub>0) \<star> inv \<phi>) \<cdot> (\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>)"
proof -
have "(s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] = \<a>[s, s\<^sub>0, w] \<cdot> ((s \<star> s\<^sub>0) \<star> inv \<phi>)"
using assms \<phi> assoc_naturality [of s s\<^sub>0 "inv \<phi>"] w.w_simps(4)
by (metis \<sigma>.leg0_simps(2-5) \<sigma>.base_simps(2-4) arr_inv cod_inv dom_inv
in_homE trg_cod)
thus ?thesis using comp_assoc by metis
qed
also have "... = \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>) \<cdot> (s\<^sub>1 \<star> \<phi>)"
proof -
have "((s \<star> s\<^sub>0) \<star> inv \<phi>) \<cdot> (\<sigma> \<star> w') = (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>)"
using \<phi> comp_arr_dom comp_cod_arr in_hhom_def
interchange [of "s \<star> s\<^sub>0" \<sigma> "inv \<phi>" w']
interchange [of \<sigma> s\<^sub>1 w "inv \<phi>"]
by auto
thus ?thesis
using comp_assoc by metis
qed
also have "... = \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w)"
proof -
have "(\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>) \<cdot> (s\<^sub>1 \<star> \<phi>) = \<sigma> \<star> w"
proof -
have "(\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>) \<cdot> (s\<^sub>1 \<star> \<phi>) = (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi> \<cdot> \<phi>)"
using \<phi> whisker_left in_hhom_def by auto
also have "... = (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> w)"
using \<phi> comp_inv_arr' by auto
also have "... = \<sigma> \<star> w"
using whisker_right [of w \<sigma> 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 "... = \<Delta>"
using assms(1) comp_assoc w.is_ide w.the_\<nu>_props(1) w.the_\<theta>_props(1) by simp
finally show ?thesis
using comp_assoc by auto
qed
qed
qed
end
text \<open>
In the special case that \<open>\<mu>\<close> is an identity 2-cell, the induced map from the apex of \<open>\<rho>\<close>
to the apex of \<open>\<sigma>\<close> is an equivalence map.
\<close>
locale identity_arrow_of_tabulations_in_maps =
arrow_of_tabulations_in_maps +
assumes is_ide: "ide \<mu>"
begin
lemma r_eq_s:
shows "r = s"
using is_ide by (metis ide_char in_hom in_homE)
lemma \<Delta>_eq_\<rho>:
shows "\<Delta> = \<rho>"
by (meson \<Delta>_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' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'
where e: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src s\<^sub>0 \<rightarrow> src r\<^sub>0\<guillemotright> \<and> \<guillemotleft>w' : src r\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright> \<and>
\<guillemotleft>\<theta> : r\<^sub>0 \<star> w \<Rightarrow> s\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : s\<^sub>1 \<Rightarrow> r\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<sigma> = (s \<star> \<theta>) \<cdot> \<a>[s, r\<^sub>0, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu>' : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
\<rho> = (s \<star> \<theta>') \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot> (\<sigma> \<star> w') \<cdot> \<nu>'"
using r_eq_s \<sigma>.apex_unique_up_to_equivalence [of \<rho> r\<^sub>0 r\<^sub>1] \<rho>.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 \<Delta>_eq_\<rho> 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 \<open>
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 \<open>\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>\<close>.
In this case, we can recover \<open>\<mu>\<close> from \<open>\<Delta>\<close> via adjoint transpose.
\<close>
locale arrow_in_bicategory_of_spans =
bicategory_of_spans V H \<a> \<i> src trg +
r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r +
s: identity_in_bicategory_of_spans V H \<a> \<i> src trg s
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and s :: 'a
and \<mu> :: 'a +
assumes in_hom: "\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>"
begin
abbreviation (input) r\<^sub>0 where "r\<^sub>0 \<equiv> tab\<^sub>0 r"
abbreviation (input) r\<^sub>1 where "r\<^sub>1 \<equiv> tab\<^sub>1 r"
abbreviation (input) s\<^sub>0 where "s\<^sub>0 \<equiv> tab\<^sub>0 s"
abbreviation (input) s\<^sub>1 where "s\<^sub>1 \<equiv> tab\<^sub>1 s"
lemma is_arrow_of_tabulations_in_maps:
shows "arrow_of_tabulations_in_maps V H \<a> \<i> src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \<mu>"
using in_hom by unfold_locales auto
end
sublocale identity_in_bicategory_of_spans \<subseteq> arrow_in_bicategory_of_spans V H \<a> \<i> 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 \<a> \<i> src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \<mu>
using is_arrow_of_tabulations_in_maps by simp
interpretation r: arrow_of_tabulations_in_maps V H \<a> \<i> 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 \<mu>_in_terms_of_\<Delta>:
shows "\<mu> = r.T0.trnr\<^sub>\<epsilon> (cod \<mu>) \<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab)"
proof -
have \<mu>: "arr \<mu>"
using in_hom by auto
have "\<mu> \<cdot> r.T0.trnr\<^sub>\<epsilon> r r.tab = r.T0.trnr\<^sub>\<epsilon> s \<Delta>"
proof -
have "\<mu> \<cdot> r.T0.trnr\<^sub>\<epsilon> r r.tab =
(\<mu> \<cdot> \<r>[r]) \<cdot> (r \<star> r.\<epsilon>) \<cdot> \<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
unfolding r.T0.trnr\<^sub>\<epsilon>_def using comp_assoc by presburger
also have "... = \<r>[s] \<cdot> ((\<mu> \<star> src \<mu>) \<cdot> (r \<star> r.\<epsilon>)) \<cdot>
\<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
using \<mu> runit_naturality comp_assoc
by (metis in_hom in_homE)
also have "... = \<r>[s] \<cdot> (s \<star> r.\<epsilon>) \<cdot> ((\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*) \<cdot>
\<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*]) \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have "(\<mu> \<star> src \<mu>) \<cdot> (r \<star> r.\<epsilon>) = \<mu> \<star> r.\<epsilon>"
using \<mu> 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 \<star> r.\<epsilon>) \<cdot> (\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*)"
using in_hom interchange [of s \<mu> r.\<epsilon> "tab\<^sub>0 r \<star> (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 "(\<mu> \<star> src \<mu>) \<cdot> (r \<star> r.\<epsilon>) = (s \<star> r.\<epsilon>) \<cdot> (\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*)"
by blast
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<r>[s] \<cdot> (s \<star> r.\<epsilon>) \<cdot> \<a>[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot>
((\<mu> \<star> tab\<^sub>0 r) \<star> (tab\<^sub>0 r)\<^sup>*) \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have "(\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*) \<cdot> \<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] =
\<a>[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot> ((\<mu> \<star> tab\<^sub>0 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
using \<mu> assoc_naturality [of \<mu> "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 "... = \<r>[s] \<cdot> (s \<star> r.\<epsilon>) \<cdot> \<a>[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot>
((\<mu> \<star> tab\<^sub>0 r) \<cdot> r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
using \<mu> whisker_right \<Delta>_simps(1) by auto
also have "... = r.T0.trnr\<^sub>\<epsilon> s \<Delta>"
unfolding r.T0.trnr\<^sub>\<epsilon>_def by simp
finally show ?thesis by blast
qed
thus ?thesis
using \<mu> 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 \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 +
\<tau>: tabulation_in_maps V H \<a> \<i> src trg t \<tau> t\<^sub>0 t\<^sub>1 +
\<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 \<mu> +
\<pi>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 \<pi>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and t :: 'a
and \<tau> :: 'a
and t\<^sub>0 :: 'a
and t\<^sub>1 :: 'a
and \<mu> :: 'a
and \<pi> :: 'a
begin
text \<open>
$$
\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]
}
$$
\<close>
notation isomorphic (infix "\<cong>" 50)
interpretation arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 \<open>\<pi> \<cdot> \<mu>\<close>
using \<mu>.in_hom \<pi>.in_hom by (unfold_locales, blast)
lemma is_arrow_of_tabulations_in_maps:
shows "arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 (\<pi> \<cdot> \<mu>)"
..
lemma chine_char:
shows "chine \<cong> \<pi>.chine \<star> \<mu>.chine"
proof -
have "is_induced_map (\<pi>.chine \<star> \<mu>.chine)"
proof -
let ?f = "\<mu>.chine"
have f: "\<guillemotleft>?f : src \<rho> \<rightarrow> src \<sigma>\<guillemotright> \<and> is_left_adjoint ?f \<and> ide ?f \<and> \<mu>.is_induced_map ?f"
using \<mu>.chine_is_induced_map \<mu>.is_map by auto
let ?g = "\<pi>.chine"
have g: "\<guillemotleft>?g : src \<sigma> \<rightarrow> src \<tau>\<guillemotright> \<and> is_left_adjoint ?g \<and> ide ?g \<and> \<pi>.is_induced_map ?g"
using \<pi>.chine_is_induced_map \<pi>.is_map by auto
let ?\<theta> = "\<mu>.the_\<theta> \<cdot> (\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]"
let ?\<nu> = "\<a>[t\<^sub>1, ?g, ?f] \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
have \<theta>: "\<guillemotleft>?\<theta> : t\<^sub>0 \<star> ?g \<star> ?f \<Rightarrow> r\<^sub>0\<guillemotright>"
using f g \<pi>.the_\<theta>_props \<mu>.the_\<theta>_props
by (intro comp_in_homI hcomp_in_vhom, auto+)
have \<nu>: "\<guillemotleft>?\<nu> : r\<^sub>1 \<Rightarrow> t\<^sub>1 \<star> ?g \<star> ?f\<guillemotright>"
using f g \<pi>.the_\<theta>_props \<mu>.the_\<theta>_props
by (intro comp_in_homI hcomp_in_vhom, auto)
interpret gf: arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 t\<^sub>0 t\<^sub>1 \<open>?g \<star> ?f\<close>
proof
show "ide (?g \<star> ?f)" by simp
show "\<exists>\<theta>. \<guillemotleft>\<theta> : t\<^sub>0 \<star> ?g \<star> ?f \<Rightarrow> r\<^sub>0\<guillemotright>"
using \<theta> by blast
show "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> t\<^sub>1 \<star> ?g \<star> ?f\<guillemotright> \<and> iso \<nu>"
using \<nu> \<mu>.the_\<nu>_props \<mu>.the_\<theta>_props \<pi>.the_\<nu>_props \<pi>.the_\<theta>_props
isos_compose [of "\<mu>.the_\<nu>" "\<pi>.the_\<nu>"] \<mu>.is_ide \<nu> \<open>ide (\<pi>.chine \<star> \<mu>.chine)\<close>
\<pi>.uw\<theta> \<pi>.w_simps(4) \<tau>.ide_leg1 \<tau>.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 \<theta>_eq: "?\<theta> = gf.the_\<theta>"
using \<theta> gf.the_\<theta>_props gf.leg0_uniquely_isomorphic by auto
have \<nu>_eq: "?\<nu> = gf.the_\<nu>"
using \<nu> gf.the_\<nu>_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 \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) t\<^sub>0 t\<^sub>1 (?g \<star> ?f)"
using gf.arrow_of_spans_of_maps_axioms by simp
have "((t \<star> gf.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f)) \<cdot> gf.the_\<nu> = \<Delta>"
proof -
have "\<Delta> = (\<pi> \<star> r\<^sub>0) \<cdot> (\<mu> \<star> r\<^sub>0) \<cdot> \<rho>"
using whisker_right comp_assoc
by (metis \<Delta>_simps(1) hseqE ide_u seqE)
also have "... = ((\<pi> \<star> r\<^sub>0) \<cdot> (s \<star> \<mu>.the_\<theta>)) \<cdot> \<a>[s, s\<^sub>0, ?f] \<cdot> (\<sigma> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using \<mu>.\<Delta>_naturality comp_assoc by simp
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> ((\<pi> \<star> s\<^sub>0 \<star> ?f) \<cdot> \<a>[s, s\<^sub>0, ?f]) \<cdot> (\<sigma> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
proof -
have "(\<pi> \<star> r\<^sub>0) \<cdot> (s \<star> \<mu>.the_\<theta>) = \<pi> \<star> \<mu>.the_\<theta>"
using f comp_arr_dom comp_cod_arr \<mu>.the_\<theta>_props \<pi>.in_hom
interchange [of \<pi> s r\<^sub>0 \<mu>.the_\<theta>]
by (metis in_homE)
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (\<pi> \<star> s\<^sub>0 \<star> ?f)"
using f comp_arr_dom comp_cod_arr \<mu>.the_\<theta>_props \<pi>.in_hom
interchange [of t \<pi> \<mu>.the_\<theta> "s\<^sub>0 \<star> ?f"]
by (metis in_homE)
finally have "(\<pi> \<star> r\<^sub>0) \<cdot> (s \<star> \<mu>.the_\<theta>) = (t \<star> \<mu>.the_\<theta>) \<cdot> (\<pi> \<star> s\<^sub>0 \<star> ?f)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot> (((\<pi> \<star> s\<^sub>0) \<star> ?f) \<cdot> (\<sigma> \<star> ?f)) \<cdot> \<mu>.the_\<nu>"
proof -
have "(\<pi> \<star> s\<^sub>0 \<star> ?f) \<cdot> \<a>[s, s\<^sub>0, ?f] = \<a>[t, s\<^sub>0, ?f] \<cdot> ((\<pi> \<star> s\<^sub>0) \<star> ?f)"
using f assoc_naturality [of \<pi> s\<^sub>0 ?f] \<pi>.in_hom by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot> (\<pi>.\<Delta> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using whisker_right comp_assoc by simp
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot>
((t \<star> \<pi>.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g] \<cdot> (\<tau> \<star> ?g) \<cdot> \<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using \<pi>.\<Delta>_naturality by simp
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot>
(((t \<star> \<pi>.the_\<theta>) \<star> ?f) \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f) \<cdot> ((\<tau> \<star> ?g) \<star> ?f) \<cdot>
(\<pi>.the_\<nu> \<star> ?f)) \<cdot> \<mu>.the_\<nu>"
using f g \<pi>.the_\<theta>_props \<pi>.the_\<nu>_props whisker_right
by (metis \<pi>.\<Delta>_simps(1) \<pi>.\<Delta>_naturality seqE)
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (\<a>[t, s\<^sub>0, ?f] \<cdot>
((t \<star> \<pi>.the_\<theta>) \<star> ?f)) \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f) \<cdot> ((\<tau> \<star> ?g) \<star> ?f) \<cdot>
(\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using comp_assoc by presburger
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (t \<star> \<pi>.the_\<theta> \<star> ?f) \<cdot>
(\<a>[t, t\<^sub>0 \<star> ?g, ?f] \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f)) \<cdot>
((\<tau> \<star> ?g) \<star> ?f) \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using f \<pi>.the_\<theta>_props assoc_naturality [of t "\<pi>.the_\<theta>" ?f] \<pi>.\<theta>_simps(3) comp_assoc
by auto
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (t \<star> \<pi>.the_\<theta> \<star> ?f) \<cdot>
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<a>[t \<star> t\<^sub>0, ?g, ?f] \<cdot>
((\<tau> \<star> ?g) \<star> ?f)) \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
proof -
have "seq \<a>[t, t\<^sub>0, ?g \<star> ?f] \<a>[t \<star> t\<^sub>0, ?g, ?f]"
using f g by fastforce
moreover have "inv (t \<star> \<a>[t\<^sub>0, ?g, ?f]) = t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]"
using f g by simp
moreover have "iso (t \<star> \<a>[t\<^sub>0, ?g, ?f])"
using f g by simp
have "\<a>[t, t\<^sub>0 \<star> ?g, ?f] \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f) =
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> \<a>[t \<star> t\<^sub>0, ?g, ?f]"
proof -
have "seq \<a>[t, t\<^sub>0, ?g \<star> ?f] \<a>[t \<star> t\<^sub>0, ?g, ?f]"
using f g by fastforce
moreover have "inv (t \<star> \<a>[t\<^sub>0, ?g, ?f]) = t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]"
using f g by simp
moreover have "iso (t \<star> \<a>[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 \<pi>.w_simps(4) \<tau>.ide_base \<tau>.ide_leg0 \<tau>.leg0_simps(3))
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> \<mu>.the_\<theta>) \<cdot> (t \<star> \<pi>.the_\<theta> \<star> ?f)) \<cdot>
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f) \<cdot>
\<a>[t\<^sub>1, ?g, ?f] \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using f g assoc_naturality [of \<tau> ?g ?f] comp_assoc by simp
also have "... = (t \<star> \<mu>.the_\<theta> \<cdot> (\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot>
\<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot>
(\<tau> \<star> ?g \<star> ?f) \<cdot> \<a>[t\<^sub>1, ?g, ?f] \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
proof -
have 1: "seq \<mu>.the_\<theta> ((\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])"
using \<theta>_eq by auto
hence "t \<star> (\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f] =
(t \<star> \<pi>.the_\<theta> \<star> ?f) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])"
using whisker_left \<tau>.ide_base by blast
thus ?thesis
using 1 whisker_left \<tau>.ide_base comp_assoc by presburger
qed
also have "... = ((t \<star> gf.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f)) \<cdot> gf.the_\<nu>"
using \<theta>_eq \<nu>_eq by (simp add: comp_assoc)
finally show ?thesis
using comp_assoc by presburger
qed
thus "((t \<star> gf.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f)) \<cdot>
arrow_of_spans_of_maps.the_\<nu> (\<cdot>) (\<star>) (dom ((\<pi> \<cdot> \<mu> \<star> r\<^sub>0) \<cdot> \<rho>)) t\<^sub>1 (?g \<star> ?f) =
\<Delta>"
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 \<subseteq>
arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 "\<pi> \<cdot> \<mu>"
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 \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 +
\<tau>: tabulation_in_maps V H \<a> \<i> src trg t \<tau> t\<^sub>0 t\<^sub>1 +
\<mu>: tabulation_in_maps V H \<a> \<i> src trg u \<mu> u\<^sub>0 u\<^sub>1 +
\<rho>\<sigma>: composite_tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 +
\<tau>\<mu>: composite_tabulation_in_maps V H \<a> \<i> src trg t \<tau> t\<^sub>0 t\<^sub>1 u \<mu> u\<^sub>0 u\<^sub>1 +
\<omega>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 \<omega> +
\<chi>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 u \<mu> u\<^sub>0 u\<^sub>1 \<chi>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and t :: 'a
and \<tau> :: 'a
and t\<^sub>0 :: 'a
and t\<^sub>1 :: 'a
and u :: 'a
and \<mu> :: 'a
and u\<^sub>0 :: 'a
and u\<^sub>1 :: 'a
and \<omega> :: 'a
and \<chi> :: 'a
begin
text \<open>
$$
\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] \\
}
$$
\<close>
notation isomorphic (infix "\<cong>" 50)
interpretation arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> \<rho>\<sigma>.tab \<open>s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0\<close> \<open>r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1\<close>
\<open>t \<star> u\<close> \<tau>\<mu>.tab \<open>u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0\<close> \<open>t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1\<close> \<open>\<omega> \<star> \<chi>\<close>
using \<rho>\<sigma>.composable \<omega>.in_hom \<chi>.in_hom
by unfold_locales auto
lemma is_arrow_of_tabulations_in_maps:
shows "arrow_of_tabulations_in_maps V H \<a> \<i> src trg
(r \<star> s) \<rho>\<sigma>.tab (s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) (r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1)
(t \<star> u) \<tau>\<mu>.tab (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) (\<omega> \<star> \<chi>)"
..
sublocale arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> \<rho>\<sigma>.tab \<open>s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0\<close> \<open>r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1\<close>
\<open>t \<star> u\<close> \<tau>\<mu>.tab \<open>u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0\<close> \<open>t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1\<close> \<open>\<omega> \<star> \<chi>\<close>
using is_arrow_of_tabulations_in_maps by simp
interpretation Maps: maps_category V H \<a> \<i> src trg ..
notation Maps.comp (infixr "\<odot>" 55)
interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg s\<^sub>1 r\<^sub>0
using \<rho>.leg0_is_map \<sigma>.leg1_is_map \<rho>\<sigma>.composable apply unfold_locales by auto
interpretation r\<^sub>0s\<^sub>1: arrow_of_tabulations_in_maps \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
\<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close> r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1
\<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close> r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1
\<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close>
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 \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg u\<^sub>1 t\<^sub>0
using \<tau>.leg0_is_map \<mu>.leg1_is_map \<tau>\<mu>.composable apply unfold_locales by auto
interpretation t\<^sub>0u\<^sub>1: arrow_of_tabulations_in_maps \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
\<open>t\<^sub>0\<^sup>* \<star> u\<^sub>1\<close> t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1
\<open>t\<^sub>0\<^sup>* \<star> u\<^sub>1\<close> t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1
\<open>t\<^sub>0\<^sup>* \<star> u\<^sub>1\<close>
using t\<^sub>0u\<^sub>1.is_arrow_of_tabulations_in_maps by simp
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text \<open>
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 \<open>T2\<close> 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 \<open>\<beta>\<close> required in the
hypotheses of \<open>T2\<close> depends, for the ``input'' tabulation, on the isomorphism constructed
for the ``output'' tabulation.
\<close>
lemma prj_chine:
shows "\<tau>\<mu>.p\<^sub>0 \<star> chine \<cong> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
and "\<tau>\<mu>.p\<^sub>1 \<star> chine \<cong> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
proof -
have 1: "arrow_of_spans_of_maps V H \<a> \<i> src trg
(s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) (r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1) (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) chine \<and>
(((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>[t \<star> u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<tau>\<mu>.tab \<star> chine)) \<cdot> the_\<nu> =
((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<rho>\<sigma>.tab"
using chine_is_induced_map by simp
let ?u\<^sub>\<tau> = "u \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0"
let ?w\<^sub>\<tau> = "\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
let ?w\<^sub>\<tau>' = "\<tau>\<mu>.p\<^sub>1 \<star> chine"
have u\<^sub>\<tau>: "ide ?u\<^sub>\<tau>"
using \<chi>.u_simps(3) by auto
have w\<^sub>\<tau>: "ide ?w\<^sub>\<tau> \<and> is_left_adjoint ?w\<^sub>\<tau>"
by (simp add: \<omega>.is_map \<rho>.T0.antipar(1) left_adjoints_compose)
have w\<^sub>\<tau>': "ide ?w\<^sub>\<tau>' \<and> is_left_adjoint ?w\<^sub>\<tau>'"
by (simp add: is_map left_adjoints_compose)
let ?\<theta>\<^sub>\<tau> = "\<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot> (\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
let ?\<theta>\<^sub>\<tau>' = "(u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]"
let ?\<beta>\<^sub>\<tau> = "\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
have \<theta>\<^sub>\<tau>: "\<guillemotleft>?\<theta>\<^sub>\<tau> : t\<^sub>0 \<star> ?w\<^sub>\<tau> \<Rightarrow> ?u\<^sub>\<tau>\<guillemotright>"
using \<rho>.T0.antipar(1) \<omega>.the_\<theta>_in_hom \<chi>.u_simps(3)
by (intro comp_in_homI, auto)
have \<theta>\<^sub>\<tau>': "\<guillemotleft>?\<theta>\<^sub>\<tau>' : t\<^sub>0 \<star> ?w\<^sub>\<tau>' \<Rightarrow> ?u\<^sub>\<tau>\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] : t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine \<Rightarrow> (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine\<guillemotright>"
using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc'_in_hom by simp
show "\<guillemotleft>t\<^sub>0u\<^sub>1.\<phi> \<star> chine : (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine \<Rightarrow> (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
using \<tau>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>(\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine : (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> ((u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine : ((u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> (u \<star> u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
using assoc_in_hom by auto
show "\<guillemotleft>\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] : (u \<star> u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> u \<star> (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
by auto
show "\<guillemotleft>u \<star> the_\<theta> : u \<star> (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> ?u\<^sub>\<tau>\<guillemotright>"
by (intro hcomp_in_vhom, auto)
qed
have \<beta>\<^sub>\<tau>: "\<guillemotleft>?\<beta>\<^sub>\<tau> : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t\<^sub>1 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> (t\<^sub>1 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1 : (t\<^sub>1 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<omega>.the_\<nu>_props \<rho>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>the_\<nu> : r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine\<guillemotright>"
using the_\<nu>_in_hom(2) by simp
show "\<guillemotleft>\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] : (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine \<Rightarrow> t\<^sub>1 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc_in_hom by simp
qed
define LHS where "LHS = (t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>)"
have LHS: "\<guillemotleft>LHS : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
proof (unfold LHS_def, intro comp_in_homI)
show "\<guillemotleft>\<tau> \<star> ?w\<^sub>\<tau> : t\<^sub>1 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> (t \<star> t\<^sub>0) \<star> ?w\<^sub>\<tau>\<guillemotright>"
using \<rho>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] : (t \<star> t\<^sub>0) \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau>\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>t \<star> ?\<theta>\<^sub>\<tau> : t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
proof -
have "src t = trg (t\<^sub>0 \<star> \<omega>.chine \<star> r\<^sub>0s\<^sub>1.p\<^sub>1)"
by (metis \<chi>.u_simps(3) \<mu>.ide_base \<sigma>.ide_leg0 \<sigma>.leg1_simps(3) \<tau>\<mu>.composable
\<theta>\<^sub>\<tau> 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 \<theta>\<^sub>\<tau> by blast
qed
qed
define RHS where "RHS = ((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) \<cdot> ?\<beta>\<^sub>\<tau>"
have RHS: "\<guillemotleft>RHS : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
unfolding RHS_def
proof
show "\<guillemotleft>?\<beta>\<^sub>\<tau> : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t\<^sub>1 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
using \<beta>\<^sub>\<tau> by simp
show "\<guillemotleft>(t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') : t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
proof
show "\<guillemotleft>\<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') : t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<Rightarrow> t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
using \<tau>.T0.antipar(1) by fastforce
show "\<guillemotleft>t \<star> ?\<theta>\<^sub>\<tau>' : t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau>' \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
using w\<^sub>\<tau>' \<theta>\<^sub>\<tau>' \<tau>.leg0_simps(2) \<tau>.leg0_simps(3) hseqI' ideD(1) t\<^sub>0u\<^sub>1.p\<^sub>1_simps
trg_hcomp \<tau>.base_in_hom(2) hcomp_in_vhom
by presburger
qed
qed
have eq: "LHS = RHS"
proof -
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) = \<Delta>"
proof -
text \<open>
Here we use \<open>\<omega>.\<Delta>_naturality\<close> to replace @{term \<omega>.chine}
in favor of @{term \<omega>}.
We have to bring @{term \<omega>.the_\<nu>}, @{term \<tau>}, and @{term \<omega>.the_\<theta>} together,
with @{term \<rho>\<sigma>.p\<^sub>1} on the right.
\<close>
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<tau> \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot>
(\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
unfolding LHS_def
using comp_assoc by presburger
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<tau> \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
\<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
(t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1])"
using whisker_left \<tau>.ide_base \<theta>\<^sub>\<tau> arrI seqE
by (metis (full_types))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> ((t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> \<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(\<tau> \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
\<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1)"
using assoc_naturality
by (metis \<omega>.w_simps(2-6) \<rho>.leg1_simps(3) \<rho>\<sigma>.leg1_simps(2) \<tau>.tab_simps(1)
\<tau>.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 "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
((t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
(\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> ((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
(\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> \<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot>
\<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
\<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "seq \<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
by (simp add: \<rho>.T0.antipar(1))
moreover have "inv (t \<star> \<a>[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) = t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using \<rho>.T0.antipar(1) by simp
ultimately show ?thesis
using pentagon \<rho>.T0.antipar(1) iso_hcomp
invert_side_of_triangle(1)
[of "\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> \<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
"t \<star> \<a>[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
"\<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1)"]
by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (((t \<star> \<omega>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
(\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> ((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1)) \<cdot>
(\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((t \<star> \<omega>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>1)"
using assoc_naturality [of t \<omega>.the_\<theta> \<rho>\<sigma>.p\<^sub>1] \<omega>.\<theta>_simps(3) \<rho>\<sigma>.leg1_simps(2) hseq_char
by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<cdot> \<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_right \<rho>.T0.antipar(1) \<omega>.\<Delta>_simps(1) \<omega>.\<Delta>_naturality comp_assoc
by fastforce
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> ((t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<cdot> \<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0 = (t \<star> (\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)"
using whisker_left whisker_right \<rho>.T0.antipar(1)
by (metis (full_types) \<chi>.\<Delta>_simps(1) \<tau>.ide_base \<theta>\<^sub>\<tau> arrI r\<^sub>0s\<^sub>1.ide_u seqE)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<cdot> \<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> (\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0) =
t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0)"
using \<chi>.in_hom whisker_left by auto
also have "... = t \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]"
using assoc_naturality [of \<chi> s\<^sub>0 \<rho>\<sigma>.p\<^sub>0] \<chi>.in_hom by auto
also have "... = (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0])"
proof -
have "seq (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]"
using \<chi>.in_hom
apply (intro seqI hseqI)
apply auto
proof -
show "\<guillemotleft>\<chi> : src u \<rightarrow> trg \<chi>\<guillemotright>"
by (metis \<chi>.\<Delta>_simps(1) \<chi>.u_simps(3) hseqE in_hhom_def seqE)
show "dom (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) = s \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0"
by (metis \<Delta>_simps(1) \<chi>.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 "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>1)) \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_right comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(\<omega> \<star> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1)) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using assoc_naturality [of \<omega> r\<^sub>0 \<rho>\<sigma>.p\<^sub>1] \<omega>.in_hom \<rho>.T0.antipar(1) comp_assoc
by fastforce
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> ((t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> (\<omega> \<star> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1) = \<omega> \<star> r\<^sub>0s\<^sub>1.\<phi>"
using comp_cod_arr comp_arr_dom \<omega>.in_hom interchange comp_ide_arr
by (metis \<tau>.base_in_hom(2) \<tau>.ide_base r\<^sub>0s\<^sub>1.\<phi>_simps(1) r\<^sub>0s\<^sub>1.\<phi>_simps(4) seqI')
also have "... = (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi>)"
using r\<^sub>0s\<^sub>1.\<phi>_in_hom comp_cod_arr comp_arr_dom \<omega>.in_hom interchange
by (metis in_homE)
finally have "(t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> (\<omega> \<star> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1) = (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi>)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
((t \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_left \<rho>.T0.antipar(1) \<rho>\<sigma>.composable \<chi>.in_hom comp_assoc by auto
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<omega> \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) =
\<omega> \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)"
proof -
have "\<guillemotleft>(\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) : s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0 \<Rightarrow> u \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
using \<omega>.in_hom \<chi>.in_hom by force
thus ?thesis
by (metis (no_types) \<omega>.in_hom comp_arr_dom comp_cod_arr in_homE
interchange)
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<omega> \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (r \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "\<omega> \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) =
(\<omega> \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (r \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0))"
proof -
have "seq (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) (\<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0))"
using \<chi>.in_hom by force
thus ?thesis
using comp_arr_dom comp_cod_arr \<omega>.in_hom \<chi>.in_hom interchange
by (metis in_homE)
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<omega> \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) =
((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0]"
using assoc_naturality \<omega>.in_hom \<chi>.in_hom
by (metis \<rho>\<sigma>.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 "... = \<Delta>"
using whisker_left \<rho>\<sigma>.tab_def comp_assoc by simp
finally show ?thesis by auto
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
text \<open>Now cancel @{term \<omega>.the_\<nu>} and its inverse.\<close>
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
(\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> (\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot>
\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
((\<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
unfolding RHS_def
using comp_assoc by presburger
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
(\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> (\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot>
\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu>"
proof -
have "the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
((\<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
((t\<^sub>1 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
using comp_inv_arr \<rho>.T0.antipar(1) comp_assoc_assoc' by simp
also have "... = the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
using comp_cod_arr \<rho>.T0.antipar(1) by simp
also have "... = the_\<nu> \<cdot> (r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_right [of \<rho>\<sigma>.p\<^sub>1] r\<^sub>0s\<^sub>1.ide_leg1 \<omega>.the_\<nu>_props(2) \<omega>.the_\<nu>_simps(4)
\<rho>.leg1_simps(2) comp_inv_arr'
by metis
also have "... = the_\<nu>"
using comp_arr_dom by simp
finally show ?thesis
using comp_assoc by simp
qed
text \<open>
Now reassociate to move @{term the_\<theta>} to the left and get other terms composed
with @{term chine}, where they can be reduced to @{term \<tau>\<mu>.tab}.
\<close>
also have "... = (\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> u \<star> the_\<theta>)) \<cdot> (t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
(t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> (t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> (\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot>
\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu>"
proof -
have "arr ((u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using \<theta>\<^sub>\<tau>' by blast
moreover have "arr (\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
moreover have "arr ((\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
moreover have "arr (((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
moreover have "arr ((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
ultimately
have "t \<star> (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] =
(t \<star> u \<star> the_\<theta>) \<cdot> (t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
(t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> (t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using whisker_left \<rho>.T0.antipar(1) \<tau>.ide_base by presburger
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot>
((\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot> \<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> the_\<nu>"
using assoc'_naturality [of t u the_\<theta>] \<tau>\<mu>.composable \<theta>_simps(3) comp_assoc by auto
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot>
\<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot> \<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] =
\<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine)"
using assoc_naturality
by (metis \<tau>.leg1_simps(3) \<tau>.tab_simps(1,2,4) \<tau>.tab_simps(5) \<tau>\<mu>.leg0_simps(2)
\<tau>\<mu>.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 \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
((t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot>
\<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] =
\<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine)"
using pentagon t\<^sub>0u\<^sub>1.p\<^sub>1_simps uw\<theta> \<tau>.T0.antipar(1) iso_hcomp
comp_assoc_assoc'
invert_side_of_triangle(1)
[of "\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> \<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]"
"t \<star> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]"
"\<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine)"]
by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>[t, u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine] =
\<a>[t, u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine)"
using assoc_naturality [of t t\<^sub>0u\<^sub>1.\<phi> chine] t\<^sub>0u\<^sub>1.cospan by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>[t, u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0, chine] =
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> ((t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0)) \<star> chine)"
using assoc_naturality [of t "\<mu> \<star> \<tau>\<mu>.p\<^sub>0" chine]
by (simp add: \<tau>\<mu>.composable)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) =
((t \<star> ((u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0)) \<star> chine) \<cdot> ((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine)"
using whisker_right whisker_left [of t "\<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]" "\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]"]
\<tau>\<mu>.composable comp_assoc_assoc'
by simp
also have "... = (t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine"
using comp_cod_arr \<tau>\<mu>.composable by simp
finally have "(((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) =
(t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine"
by simp
thus ?thesis
using comp_assoc by metis
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (((\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine)) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "((\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> (\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine)) \<cdot>
((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) =
((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)"
using comp_inv_arr' comp_cod_arr \<tau>\<mu>.composable comp_assoc_assoc'
whisker_right [of chine "\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0]" "\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine)) \<cdot> the_\<nu>"
using comp_assoc by presburger
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot>
(\<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine)) \<cdot>
(\<tau>\<mu>.tab \<star> chine) \<cdot> the_\<nu>"
proof -
have "(\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) =
\<tau>\<mu>.tab \<star> chine"
using uw\<theta> whisker_right [of chine]
by (metis \<tau>\<mu>.tab_def \<tau>\<mu>.tab_in_vhom' arrI seqE)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>[t \<star> u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<tau>\<mu>.tab \<star> chine) \<cdot> the_\<nu>"
proof -
have "\<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot> (t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
(t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> \<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle>, \<^bold>\<langle>u\<^bold>\<rangle>, (\<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle>, (\<^bold>\<langle>u\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle>, \<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>)\<rbrace>"
using \<a>'_def \<alpha>_def \<tau>\<mu>.composable by simp
also have "... = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<tau>\<mu>.composable
apply (intro E.eval_eqI) by simp_all
also have "... = \<a>[t \<star> u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]"
using \<a>'_def \<alpha>_def \<tau>\<mu>.composable by simp
finally show ?thesis by simp
qed
also have "... = \<Delta>"
using \<Delta>_naturality by simp
finally show ?thesis by simp
qed
finally have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.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 \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
(RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
using u\<^sub>\<tau> r\<^sub>0s\<^sub>1.ide_u LHS RHS iso_is_section [of "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0]"] section_is_mono
monoE \<tau>\<mu>.composable comp_assoc
by (metis (no_types, lifting) \<Delta>_simps(1) \<mu>.ide_base
\<open>\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> r\<^sub>0s\<^sub>1.p\<^sub>1) =
((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0) \<cdot> \<rho>\<sigma>.tab\<close>
\<tau>.ide_base hseq_char ideD(1) ide_u iso_assoc')
hence 1: "LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] = RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using epiE LHS RHS iso_is_retraction retraction_is_epi \<tau>\<mu>.composable
\<omega>.the_\<nu>_props iso_hcomp
by (metis \<Delta>_simps(1) \<omega>.the_\<nu>_simps(2)
\<open>((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0) \<cdot> \<rho>\<sigma>.tab =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> r\<^sub>0s\<^sub>1.p\<^sub>1)\<close>
\<rho>.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 \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using iso_is_retraction retraction_is_epi \<rho>.T0.antipar(1) by simp
moreover have "seq LHS \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using LHS \<rho>.T0.antipar(1) by auto
moreover have "seq RHS \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using RHS \<rho>.T0.antipar(1) by auto
ultimately show ?thesis
using epiE 1 by blast
qed
qed
have 1: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)"
using LHS_def RHS_def u\<^sub>\<tau> w\<^sub>\<tau> w\<^sub>\<tau>' \<beta>\<^sub>\<tau> \<theta>\<^sub>\<tau> \<theta>\<^sub>\<tau>' eq \<tau>.T2 [of ?w\<^sub>\<tau> ?w\<^sub>\<tau>' ?\<theta>\<^sub>\<tau> ?u\<^sub>\<tau> ?\<theta>\<^sub>\<tau>' ?\<beta>\<^sub>\<tau>]
by fastforce
obtain \<gamma>\<^sub>\<tau> where \<gamma>\<^sub>\<tau>: "\<guillemotleft>\<gamma>\<^sub>\<tau> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma>\<^sub>\<tau> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>)"
using 1 by auto
text \<open>
At this point we could show that @{term \<gamma>\<^sub>\<tau>} is invertible using \<open>BS3\<close>,
but we want to avoid using \<open>BS3\<close> if possible and we also want to
establish a characterization of @{term "inv \<gamma>\<^sub>\<tau>"}. So we show the invertibility of
@{term \<gamma>\<^sub>\<tau>} directly, using a few more applications of \<open>T2\<close>.
\<close>
have iso_\<beta>\<^sub>\<tau>: "iso ?\<beta>\<^sub>\<tau>"
using uw\<theta> \<beta>\<^sub>\<tau> the_\<nu>_props \<omega>.the_\<nu>_props hseqI' iso_assoc' \<omega>.hseq_leg\<^sub>0
apply (intro isos_compose)
apply (metis \<omega>.is_ide \<rho>\<sigma>.leg1_simps(2) \<tau>.ide_leg1 \<tau>.leg1_simps(2)
\<tau>.leg1_simps(3) hseqE r\<^sub>0s\<^sub>1.ide_leg1 hcomp_simps(1) vconn_implies_hpar(3))
apply (metis \<rho>\<sigma>.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 \<tau>.ide_leg1 \<tau>.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>\<tau>')
by blast
hence eq': "((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) =
((t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>)) \<cdot> inv ?\<beta>\<^sub>\<tau>"
proof -
have "seq ((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) ?\<beta>\<^sub>\<tau>"
using LHS RHS_def eq by blast
hence "(t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') =
(((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) \<cdot> ?\<beta>\<^sub>\<tau>) \<cdot> inv ?\<beta>\<^sub>\<tau>"
by (meson invert_side_of_triangle(2) iso_\<beta>\<^sub>\<tau>)
thus ?thesis
using LHS_def RHS_def eq by argo
qed
have 2: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> inv ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>)"
using u\<^sub>\<tau> w\<^sub>\<tau> w\<^sub>\<tau>' \<beta>\<^sub>\<tau> \<theta>\<^sub>\<tau> \<theta>\<^sub>\<tau>' eq' \<tau>.T2 [of ?w\<^sub>\<tau>' ?w\<^sub>\<tau> ?\<theta>\<^sub>\<tau>'?u\<^sub>\<tau> ?\<theta>\<^sub>\<tau> "inv ?\<beta>\<^sub>\<tau>"] iso_\<beta>\<^sub>\<tau> comp_assoc
by blast
obtain \<gamma>\<^sub>\<tau>' where
\<gamma>\<^sub>\<tau>': "\<guillemotleft>\<gamma>\<^sub>\<tau>' : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> inv ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>' \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')"
using 2 by auto
have "inverse_arrows \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>'"
proof
have "\<guillemotleft>\<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright>"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' by auto
moreover have "t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau> = t\<^sub>1 \<star> ?w\<^sub>\<tau>"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' whisker_left \<beta>\<^sub>\<tau> iso_\<beta>\<^sub>\<tau> comp_inv_arr'
by (metis (no_types, lifting) \<tau>.ide_leg1 calculation in_homE)
moreover have "?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau>)"
proof -
have "?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>)"
using \<gamma>\<^sub>\<tau> by simp
also have "... = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>') \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>)"
using \<gamma>\<^sub>\<tau>' comp_assoc by simp
also have "... = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau>)"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' whisker_left
by (metis (full_types) \<tau>.ide_leg0 seqI')
finally show ?thesis by simp
qed
moreover have
"\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>) \<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>"
proof -
have "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> t\<^sub>1 \<star> ?w\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>)"
proof -
have "(t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>) =
((t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>)) \<cdot> (t\<^sub>1 \<star> ?w\<^sub>\<tau>)"
by (metis LHS LHS_def comp_arr_dom in_homE)
thus ?thesis
using w\<^sub>\<tau> \<theta>\<^sub>\<tau> \<omega>.w_simps(4) \<tau>.leg1_in_hom(2) \<tau>.leg1_simps(3) hcomp_in_vhom ideD(1)
trg_hcomp ide_in_hom(2) \<tau>.T2
by presburger
qed
thus "\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>) \<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>"
by (metis \<theta>\<^sub>\<tau> comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\<tau>)
qed
ultimately have "\<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau> = ?w\<^sub>\<tau>"
by simp
thus "ide (\<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau>)"
using w\<^sub>\<tau> by simp
have "\<guillemotleft>\<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright>"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' by auto
moreover have "t\<^sub>1 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' = t\<^sub>1 \<star> ?w\<^sub>\<tau>'"
by (metis \<beta>\<^sub>\<tau> \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' \<tau>.ide_leg1 calculation comp_arr_inv' in_homE iso_\<beta>\<^sub>\<tau> whisker_left)
moreover have "?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>')"
proof -
have "?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')"
using \<gamma>\<^sub>\<tau>' by simp
also have "... = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>) \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')"
using \<gamma>\<^sub>\<tau> comp_assoc by simp
also have "... = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>')"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' whisker_left \<tau>.ide_leg0 by fastforce
finally show ?thesis by simp
qed
moreover have "\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)
\<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>'"
proof -
have "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> t\<^sub>1 \<star> ?w\<^sub>\<tau>' = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)"
proof -
have "(t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') =
((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) \<cdot> (t\<^sub>1 \<star> ?w\<^sub>\<tau>')"
proof -
have 1: "t\<^sub>1 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' = (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>) \<cdot> (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>')"
by (meson \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' \<tau>.ide_leg1 seqI' whisker_left)
have "((LHS \<cdot> inv ?\<beta>\<^sub>\<tau>) \<cdot> (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>)) \<cdot> (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>') = LHS \<cdot> inv ?\<beta>\<^sub>\<tau>"
using LHS_def RHS_def \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' 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>\<tau>' \<theta>\<^sub>\<tau>' \<omega>.w_simps(4) \<tau>.leg1_in_hom(2) \<tau>.leg1_simps(3) hcomp_in_vhom ideD(1)
trg_hcomp ide_in_hom(2) \<tau>.T2 \<tau>.T0.antipar(1) t\<^sub>0u\<^sub>1.base_simps(2)
t\<^sub>0u\<^sub>1.leg1_simps(4)
by presburger
qed
thus "\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)
\<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>'"
by (metis \<theta>\<^sub>\<tau>' comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\<tau>')
qed
ultimately have "\<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' = ?w\<^sub>\<tau>'"
by simp
thus "ide (\<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>')"
using w\<^sub>\<tau>' by simp
qed
thus "\<tau>\<mu>.p\<^sub>1 \<star> chine \<cong> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
using w\<^sub>\<tau> w\<^sub>\<tau>' \<gamma>\<^sub>\<tau> isomorphic_symmetric isomorphic_def by blast
have iso_\<gamma>\<^sub>\<tau>: "iso \<gamma>\<^sub>\<tau>"
using \<open>inverse_arrows \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>'\<close> by auto
have \<gamma>\<^sub>\<tau>'_eq: "\<gamma>\<^sub>\<tau>' = inv \<gamma>\<^sub>\<tau>"
using \<open>inverse_arrows \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>'\<close> inverse_unique by blast
let ?w\<^sub>\<mu> = "\<tau>\<mu>.p\<^sub>0 \<star> chine"
let ?w\<^sub>\<mu>' = "\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
let ?u\<^sub>\<mu> = "s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0"
let ?\<theta>\<^sub>\<mu> = "the_\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]"
let ?\<theta>\<^sub>\<mu>' = "(\<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
let ?\<beta>\<^sub>\<mu> = "\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot> (\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
have w\<^sub>\<mu>: "ide ?w\<^sub>\<mu> \<and> is_left_adjoint ?w\<^sub>\<mu>"
using is_map left_adjoints_compose by simp
have w\<^sub>\<mu>': "ide ?w\<^sub>\<mu>' \<and> is_left_adjoint ?w\<^sub>\<mu>'"
using \<chi>.is_map left_adjoints_compose
by (simp add: is_map left_adjoints_compose)
have 1: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<mu> \<Rightarrow> ?w\<^sub>\<mu>'\<guillemotright> \<and> ?\<beta>\<^sub>\<mu> = u\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<mu> = ?\<theta>\<^sub>\<mu>' \<cdot> (u\<^sub>0 \<star> \<gamma>)"
proof -
have \<theta>\<^sub>\<mu>: "\<guillemotleft>?\<theta>\<^sub>\<mu> : u\<^sub>0 \<star> ?w\<^sub>\<mu> \<Rightarrow> ?u\<^sub>\<mu>\<guillemotright>"
by auto
have \<theta>\<^sub>\<mu>': "\<guillemotleft>?\<theta>\<^sub>\<mu>' : u\<^sub>0 \<star> ?w\<^sub>\<mu>' \<Rightarrow> ?u\<^sub>\<mu>\<guillemotright>"
by fastforce
have \<beta>\<^sub>\<mu>: "\<guillemotleft>?\<beta>\<^sub>\<mu> : u\<^sub>1 \<star> ?w\<^sub>\<mu> \<Rightarrow> u\<^sub>1 \<star> ?w\<^sub>\<mu>'\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine] : u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0 \<star> chine \<Rightarrow> (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
by auto
show "\<guillemotleft>inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine : (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine\<guillemotright>"
using t\<^sub>0u\<^sub>1.\<phi>_uniqueness(2) hcomp_in_vhom
by (simp add: t\<^sub>0u\<^sub>1.\<phi>_in_hom(2) w_in_hom(2))
show "\<guillemotleft>\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] : (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine \<Rightarrow> t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine\<guillemotright>"
using \<tau>.T0.antipar(1) by auto
show "\<guillemotleft>t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau> : t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine \<Rightarrow> t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<gamma>\<^sub>\<tau> iso_\<gamma>\<^sub>\<tau> using \<tau>.T0.antipar(1) by auto
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] : t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> (t\<^sub>0 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1 : (t\<^sub>0 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>r\<^sub>0s\<^sub>1.\<phi> : r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
by auto
show "\<guillemotleft>\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0 : s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0 \<Rightarrow> (u\<^sub>1 \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
by auto
show "\<guillemotleft>\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] : (u\<^sub>1 \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0 \<Rightarrow> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
by auto
qed
text \<open>
The proof of the equation below needs to make use of the equation
\<open>\<theta>\<^sub>\<tau>' = \<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')\<close> from the previous section. So the overall strategy is to
work toward an expression of the form \<open>\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')\<close> and perform the replacement
to eliminate \<open>\<gamma>\<^sub>\<tau>'\<close>.
\<close>
have eq\<^sub>\<mu>: "(u \<star> ?\<theta>\<^sub>\<mu>) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>) =
((u \<star> ?\<theta>\<^sub>\<mu>') \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>'] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>')) \<cdot> ?\<beta>\<^sub>\<mu>"
proof -
let ?LHS = "(u \<star> ?\<theta>\<^sub>\<mu>) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
let ?RHS = "((u \<star> ?\<theta>\<^sub>\<mu>') \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>'] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>')) \<cdot> ?\<beta>\<^sub>\<mu>"
have "?RHS = (u \<star> (\<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
\<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by presburger
also have "... = (u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
\<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0]) \<cdot> (\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "u \<star> (\<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] =
(u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0])"
using whisker_left \<mu>.ide_base \<theta>\<^sub>\<mu>' by blast
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "seq (u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0])
(\<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0))"
by auto
moreover have "src u = trg \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
by simp
moreover have "inv (u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) = u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
by simp
moreover have "iso (u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0])"
by simp
moreover have "iso \<a>[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
by simp
ultimately have "(u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot> \<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0] =
\<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
using pentagon hseqI' comp_assoc
invert_opposite_sides_of_square
[of "u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
"\<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0)"
"\<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0]" "\<a>[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"]
inv_hcomp \<chi>.is_ide \<chi>.w_simps(3) \<chi>.w_simps(4) \<mu>.base_simps(2) \<mu>.ide_base
\<mu>.ide_leg0 \<mu>.leg0_simps(2) \<mu>.leg0_simps(3) \<sigma>.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 "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "(u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] =
\<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0)"
using assoc_naturality [of u \<chi>.the_\<theta> \<rho>\<sigma>.p\<^sub>0] \<chi>.\<theta>_simps(3) by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> \<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) =
((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
using assoc'_naturality [of \<mu> \<chi>.chine \<rho>\<sigma>.p\<^sub>0] by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
((\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> \<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by metis
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "(\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> \<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) =
\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0"
using comp_inv_arr' comp_cod_arr by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = (\<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "arr ((u \<star> \<chi>.the_\<theta>) \<cdot> \<a>[u, u\<^sub>0, \<chi>.chine] \<cdot> (\<mu> \<star> \<chi>.chine) \<cdot> \<chi>.the_\<nu>)"
using \<chi>.\<theta>_simps(3) by simp
hence "((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) =
(u \<star> \<chi>.the_\<theta>) \<cdot> \<a>[u, u\<^sub>0, \<chi>.chine] \<cdot> (\<mu> \<star> \<chi>.chine) \<cdot> \<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0"
using whisker_right by simp
also have "... = (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0"
using \<chi>.\<Delta>_naturality by simp
finally have "((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) =
(\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>)) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by presburger
also have "... = ?\<theta>\<^sub>\<tau>' \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using \<gamma>\<^sub>\<tau>' \<gamma>\<^sub>\<tau>'_eq by simp
also have "... = (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
((\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by presburger
also have "... = (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> ((\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine] =
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> ((\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine] =
((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> ((t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_inv_arr' \<tau>.T0.antipar(1) by auto
also have "... = ((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_cod_arr t\<^sub>0u\<^sub>1.\<phi>_uniqueness by simp
also have "... = (t\<^sub>0u\<^sub>1.\<phi> \<cdot> inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using whisker_right t\<^sub>0u\<^sub>1.\<phi>_uniqueness by simp
also have "... = ((u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_arr_inv' \<tau>.T0.antipar(1) t\<^sub>0u\<^sub>1.\<phi>_uniqueness by simp
also have "... = \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_cod_arr \<tau>.T0.antipar(1) by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = (u \<star> the_\<theta>) \<cdot> (\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
using assoc'_naturality [of \<mu> \<tau>\<mu>.p\<^sub>0 chine] comp_assoc by auto
also have "... = ((u \<star> the_\<theta>) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine])) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
using uw\<theta> pentagon comp_assoc
invert_opposite_sides_of_square
[of "u \<star> \<a>[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]"
"\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine)" "\<a>[u, u\<^sub>0, ?w\<^sub>\<mu>]"
"\<a>[u \<star> u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]"]
\<mu>.base_simps(2) \<mu>.ide_base \<mu>.ide_leg0 \<mu>.leg0_simps(2) assoc'_eq_inv_assoc
ide_hcomp hcomp_simps(1) t\<^sub>0u\<^sub>1.ide_u
by force
also have "... = (u \<star> the_\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
using whisker_left comp_assoc by simp
finally show ?thesis by simp
qed
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<mu> \<Rightarrow> ?w\<^sub>\<mu>'\<guillemotright> \<and> ?\<beta>\<^sub>\<mu> = u\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<mu> = ?\<theta>\<^sub>\<mu>' \<cdot> (u\<^sub>0 \<star> \<gamma>)"
using w\<^sub>\<mu> w\<^sub>\<mu>' \<theta>\<^sub>\<mu> \<theta>\<^sub>\<mu>' \<beta>\<^sub>\<mu> eq\<^sub>\<mu> \<mu>.T2 [of ?w\<^sub>\<mu> ?w\<^sub>\<mu>' ?\<theta>\<^sub>\<mu> ?u\<^sub>\<mu> ?\<theta>\<^sub>\<mu>' ?\<beta>\<^sub>\<mu>] by fast
qed
obtain \<gamma>\<^sub>\<mu> where \<gamma>\<^sub>\<mu>: "\<guillemotleft>\<gamma>\<^sub>\<mu> : ?w\<^sub>\<mu> \<Rightarrow> ?w\<^sub>\<mu>'\<guillemotright> \<and> ?\<beta>\<^sub>\<mu> = u\<^sub>1 \<star> \<gamma>\<^sub>\<mu> \<and> ?\<theta>\<^sub>\<mu> = ?\<theta>\<^sub>\<mu>' \<cdot> (u\<^sub>0 \<star> \<gamma>\<^sub>\<mu>)"
using 1 by auto
show "?w\<^sub>\<mu> \<cong> ?w\<^sub>\<mu>'"
using w\<^sub>\<mu> w\<^sub>\<mu>' \<gamma>\<^sub>\<mu> BS3 [of ?w\<^sub>\<mu> ?w\<^sub>\<mu>' \<gamma>\<^sub>\<mu> \<gamma>\<^sub>\<mu>] isomorphic_def by auto
qed
lemma comp_L:
shows "Maps.seq \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> =
Maps.MkArr (src (\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1)) (src t) (Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>)"
proof -
show "Maps.seq \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1)"
using \<omega>.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] \<tau>.leg0_is_map \<rho>\<sigma>.leg1_in_hom by auto
qed
thus "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> =
Maps.MkArr (src (\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1)) (src t) (Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>)"
using Maps.comp_char by auto
qed
lemma comp_R:
shows "Maps.seq \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> =
Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> r\<^sub>0s\<^sub>1.p\<^sub>0\<rbrakk>)"
proof -
show "Maps.seq \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0)"
using \<chi>.is_map r\<^sub>0s\<^sub>1.leg0_is_map left_adjoints_compose [of \<chi>.chine \<rho>\<sigma>.p\<^sub>0] by simp
thus ?thesis
using Maps.CLS_in_hom \<mu>.leg1_is_map
apply (intro Maps.seqI')
apply blast
using Maps.CLS_in_hom [of u\<^sub>1] \<mu>.leg1_is_map by simp
qed
thus "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> =
Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> r\<^sub>0s\<^sub>1.p\<^sub>0\<rbrakk>)"
using Maps.comp_char by auto
qed
lemma comp_L_eq_comp_R:
shows "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof (intro Maps.arr_eqI)
show "Maps.seq \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
using comp_L(1) by simp
show "Maps.seq \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using comp_R(1) by simp
show "Maps.Dom (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Dom (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>)"
by (metis (no_types, lifting) Maps.Dom.simps(1) \<omega>.w_simps(2) \<omega>.w_simps(3)
\<rho>.leg1_simps(3) \<rho>\<sigma>.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 (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Cod (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>)"
by (metis Maps.Cod.simps(1) \<tau>\<mu>.composable comp_L(2) comp_R(2))
have A: "Maps.Map (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using comp_L(1) Maps.comp_char by auto
have B: "Maps.Map (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>) = Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using comp_R(1) Maps.comp_char by auto
have C: "Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk> = Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
proof (intro Maps.Comp_eqI)
show "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<in> Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
proof (intro Maps.in_CompI)
show "is_iso_class \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using prj_chine(2) is_iso_classI isomorphic_implies_hpar(2) by blast
show "is_iso_class \<lbrakk>t\<^sub>0\<rbrakk>"
using is_iso_classI by auto
show "\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<in> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using ide_in_iso_class prj_chine(2) isomorphic_implies_hpar(2) by blast
show "t\<^sub>0 \<in> \<lbrakk>t\<^sub>0\<rbrakk>"
using ide_in_iso_class by simp
show "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<cong> t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
using isomorphic_reflexive prj_chine(2) isomorphic_implies_hpar(2) by auto
qed
show "u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0 \<in> Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
proof (intro Maps.in_CompI)
show "is_iso_class \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using is_iso_classI by simp
show "is_iso_class \<lbrakk>u\<^sub>1\<rbrakk>"
using is_iso_classI by simp
show "\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0 \<in> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using ide_in_iso_class by simp
show "u\<^sub>1 \<in> iso_class u\<^sub>1"
using ide_in_iso_class by simp
show "u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0 \<cong> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
using isomorphic_reflexive isomorphic_implies_hpar(2) by auto
qed
show "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<cong> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
proof -
have "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<cong> (t\<^sub>0 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1"
using assoc'_in_hom [of t\<^sub>0 \<omega>.chine \<rho>\<sigma>.p\<^sub>1] iso_assoc' isomorphic_def r\<^sub>0s\<^sub>1.p\<^sub>1_simps
by auto
also have "... \<cong> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1"
using \<omega>.leg0_uniquely_isomorphic hcomp_isomorphic_ide
by (simp add: \<rho>.T0.antipar(1))
also have "... \<cong> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0"
using isomorphic_def r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) by blast
also have "... \<cong> (u\<^sub>1 \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0"
using \<chi>.leg1_uniquely_isomorphic hcomp_isomorphic_ide by auto
also have "... \<cong> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
using assoc_in_hom [of u\<^sub>1 \<chi>.chine \<rho>\<sigma>.p\<^sub>0] iso_assoc isomorphic_def by auto
finally show ?thesis by simp
qed
qed
show "Maps.Map (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Map (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>)"
using A B C by simp
qed
lemma csq:
shows "Maps.commutative_square \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof
show "Maps.cospan \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> = Maps.cod \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
using comp_L(1) by auto
show "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using comp_L_eq_comp_R by simp
qed
lemma CLS_chine:
shows "\<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = Maps.tuple \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
let ?T = "Maps.tuple \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
have "\<exists>!l. \<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> l = \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<and> \<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> l = \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using csq \<tau>\<mu>.prj_char
Maps.universal [of "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"]
by simp
moreover have "\<lbrakk>\<lbrakk>\<tau>\<mu>.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> ?T = \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<and>
\<lbrakk>\<lbrakk>\<tau>\<mu>.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> ?T = \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using csq \<tau>\<mu>.prj_char
Maps.prj_tuple [of "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"]
by simp
moreover have "\<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<and>
\<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using prj_chine \<tau>\<mu>.leg0_is_map \<tau>\<mu>.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 "\<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = ?T" by auto
qed
end
subsection "Equivalence of B and Span(Maps(B))"
subsubsection "The Functor SPN"
text \<open>
We now define a function \<open>SPN\<close> on arrows and will ultimately show that it extends to a
biequivalence from the underlying bicategory \<open>B\<close> to \<open>Span(Maps(B))\<close>.
The idea is that \<open>SPN\<close> takes \<open>\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>\<close> to the isomorphism class of an induced arrow
of spans from the chosen tabulation of \<open>r\<close> to the chosen tabulation of \<open>s\<close>.
To obtain this, we first use isomorphisms \<open>r.tab\<^sub>1 \<star> r.tab\<^sub>0\<^sup>* \<cong> r\<close> and \<open>s.tab\<^sub>1 \<star> s.tab\<^sub>0\<^sup>* \<cong> s\<close>
to transform \<open>\<mu>\<close> to \<open>\<guillemotleft>\<mu>' : r.tab\<^sub>1 \<star> r.tab\<^sub>0\<^sup>* \<Rightarrow> s.tab\<^sub>1 \<star> s.tab\<^sub>0\<^sup>*\<guillemotright>\<close>.
We then take the adjoint transpose of \<open>\<mu>'\<close> to obtain
\<open>\<guillemotleft>\<omega> : r.tab\<^sub>1 \<Rightarrow> (s.tab\<^sub>1 \<star> s.tab\<^sub>0\<^sup>*) \<star> r.tab\<^sub>0\<guillemotright>\<close>. The 2-cell \<open>\<omega>\<close> induces a map \<open>w\<close>
which is an arrow of spans from \<open>(r.tab\<^sub>0, r.tab\<^sub>1)\<close> to \<open>(s.tab\<^sub>0, s.tab\<^sub>1)\<close>.
We take the arrow of \<open>Span(Maps(B))\<close> defined by \<open>w\<close> as the value of \<open>SPN \<mu>\<close>.
Ensuring that \<open>SPN\<close> 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 \<open>r\<close> in
the proper relationship with the canonical tabulation assigned to its chosen factorization
\<open>r = g \<star> f\<^sup>*\<close>.
\<close>
context bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 50)
definition spn
where "spn \<mu> \<equiv>
arrow_of_tabulations_in_maps.chine V H \<a> \<i> src trg
(tab_of_ide (dom \<mu>)) (tab\<^sub>0 (dom \<mu>)) (cod \<mu>)
(tab_of_ide (cod \<mu>)) (tab\<^sub>0 (cod \<mu>)) (tab\<^sub>1 (cod \<mu>)) \<mu>"
lemma is_induced_map_spn:
assumes "arr \<mu>"
shows "arrow_of_tabulations_in_maps.is_induced_map V H \<a> \<i> src trg
(tab_of_ide (dom \<mu>)) (tab\<^sub>0 (dom \<mu>)) (cod \<mu>)
(tab_of_ide (cod \<mu>)) (tab\<^sub>0 (cod \<mu>)) (tab\<^sub>1 (cod \<mu>))
\<mu> (spn \<mu>)"
proof -
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>
using assms by unfold_locales auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<mu>.s.tab \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
show ?thesis
unfolding spn_def
using \<mu>.chine_is_induced_map by blast
qed
lemma spn_props:
assumes "arr \<mu>"
shows "\<guillemotleft>spn \<mu> : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src (tab\<^sub>0 (cod \<mu>))\<guillemotright>"
and "is_left_adjoint (spn \<mu>)"
and "tab\<^sub>0 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>0 (dom \<mu>)"
and "tab\<^sub>1 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>1 (dom \<mu>)"
proof -
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>
using assms by unfold_locales auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<mu>.s.tab \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
show "\<guillemotleft>spn \<mu> : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src (tab\<^sub>0 (cod \<mu>))\<guillemotright>"
using spn_def by simp
show "is_left_adjoint (spn \<mu>)"
using spn_def by (simp add: \<mu>.is_map)
show "tab\<^sub>0 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>0 (dom \<mu>)"
using spn_def isomorphic_def \<mu>.leg0_uniquely_isomorphic(1) by auto
show "tab\<^sub>1 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>1 (dom \<mu>)"
using spn_def isomorphic_def isomorphic_symmetric
\<mu>.leg1_uniquely_isomorphic(1)
by auto
qed
lemma spn_in_hom [intro]:
assumes "arr \<mu>"
shows "\<guillemotleft>spn \<mu> : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src (tab\<^sub>0 (cod \<mu>))\<guillemotright>"
and "\<guillemotleft>spn \<mu> : spn \<mu> \<Rightarrow> spn \<mu>\<guillemotright>"
using assms spn_props left_adjoint_is_ide by auto
lemma spn_simps [simp]:
assumes "arr \<mu>"
shows "is_left_adjoint (spn \<mu>)"
and "ide (spn \<mu>)"
and "src (spn \<mu>) = src (tab\<^sub>0 (dom \<mu>))"
and "trg (spn \<mu>) = src (tab\<^sub>0 (cod \<mu>))"
using assms spn_props left_adjoint_is_ide by auto
text \<open>
We need the next result to show that \<open>SPN\<close> is functorial; in particular,
that it takes \<open>\<guillemotleft>r : r \<Rightarrow> r\<guillemotright>\<close> in the underlying bicategory to a 1-cell
in \<open>Span(Maps(B))\<close>. The 1-cells in \<open>Span(Maps(B))\<close> have objects of \<open>Maps(B)\<close>
as their chines, and objects of \<open>Maps(B)\<close> are isomorphism classes of objects in the
underlying bicategory \<open>B\<close>. So we need the induced map associated with \<open>r\<close> to be isomorphic
to an object.
\<close>
lemma spn_ide:
assumes "ide r"
shows "spn r \<cong> src (tab\<^sub>0 r)"
proof -
interpret r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r
using assms by unfold_locales auto
interpret r: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> r
using r.is_arrow_of_tabulations_in_maps by simp
interpret tab: tabulation V H \<a> \<i> src trg r \<open>r.tab\<close> \<open>tab\<^sub>0 r\<close> \<open>dom r.tab\<close>
using assms r.tab_is_tabulation by simp
interpret tab: tabulation_in_maps V H \<a> \<i> src trg r \<open>r.tab\<close> \<open>tab\<^sub>0 r\<close> \<open>dom r.tab\<close>
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 \<open>
The other key result we need to show that \<open>SPN\<close> is functorial is to show
that the induced map of a composite is isomorphic to the composite of
induced maps.
\<close>
lemma spn_hcomp:
assumes "seq \<tau> \<mu>" and "g \<cong> spn \<tau>" and "f \<cong> spn \<mu>"
shows "spn (\<tau> \<cdot> \<mu>) \<cong> g \<star> f"
proof -
interpret \<tau>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<tau>\<close> \<open>cod \<tau>\<close> \<tau>
using assms by unfold_locales auto
interpret \<tau>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<tau>\<close> \<tau>.r.tab \<open>tab\<^sub>0 (dom \<tau>)\<close> \<open>tab\<^sub>1 (dom \<tau>)\<close>
\<open>cod \<tau>\<close> \<tau>.s.tab \<open>tab\<^sub>0 (cod \<tau>)\<close> \<open>tab\<^sub>1 (cod \<tau>)\<close>
\<tau>
using \<tau>.is_arrow_of_tabulations_in_maps by simp
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>dom \<tau>\<close> \<mu>
using assms apply unfold_locales
apply auto[1]
by (elim seqE, auto)
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>dom \<tau>\<close> \<tau>.r.tab \<open>tab\<^sub>0 (dom \<tau>)\<close> \<open>tab\<^sub>1 (dom \<tau>)\<close>
\<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
interpret \<tau>\<mu>: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>dom \<tau>\<close> \<tau>.r.tab \<open>tab\<^sub>0 (dom \<tau>)\<close> \<open>tab\<^sub>1 (dom \<tau>)\<close>
\<open>cod \<tau>\<close> \<tau>.s.tab \<open>tab\<^sub>0 (cod \<tau>)\<close> \<open>tab\<^sub>1 (cod \<tau>)\<close>
\<mu> \<tau>
..
have "g \<cong> \<tau>.chine"
using assms(2) spn_def by auto
moreover have "f \<cong> \<mu>.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 \<mu>.chine"
using calculation(1) isomorphic_implies_hpar(3) by auto
ultimately have "g \<star> f \<cong> \<tau>.chine \<star> \<mu>.chine"
using hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_transitive
by (meson \<mu>.is_ide isomorphic_implies_ide(1))
also have "... \<cong> spn (\<tau> \<cdot> \<mu>)"
using spn_def \<tau>\<mu>.chine_char isomorphic_symmetric
by (metis \<tau>\<mu>.in_hom in_homE)
finally show ?thesis
using isomorphic_symmetric by simp
qed
abbreviation (input) SPN\<^sub>0
where "SPN\<^sub>0 r \<equiv> Span.mkIde \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>"
definition SPN
where "SPN \<mu> \<equiv> if arr \<mu> then
\<lparr>Chn = \<lbrakk>\<lbrakk>spn \<mu>\<rbrakk>\<rbrakk>,
Dom = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>\<rbrakk>\<rparr>,
Cod = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (cod \<mu>)\<rbrakk>\<rbrakk>\<rparr>\<rparr>
else Span.null"
lemma Dom_SPN [simp]:
assumes "arr \<mu>"
shows "Dom (SPN \<mu>) = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>\<rbrakk>\<rparr>"
using assms SPN_def by simp
lemma Cod_SPN [simp]:
assumes "arr \<mu>"
shows "Cod (SPN \<mu>) = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (cod \<mu>)\<rbrakk>\<rbrakk>\<rparr>"
using assms SPN_def by simp
text \<open>Now we have to show this does the right thing for us.\<close>
lemma SPN_in_hom:
assumes "arr \<mu>"
shows "Span.in_hom (SPN \<mu>) (SPN\<^sub>0 (dom \<mu>)) (SPN\<^sub>0 (cod \<mu>))"
proof
interpret Dom: span_in_category Maps.comp \<open>Dom (SPN \<mu>)\<close>
proof
interpret r: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close>
using assms by unfold_locales auto
show "Maps.span (Leg0 (Dom (SPN \<mu>))) (Leg1 (Dom (SPN \<mu>)))"
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
\<open>\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>\<rbrakk>\<rparr>\<close>
using assms Dom.span_in_category_axioms SPN_def by simp
interpret Cod: span_in_category Maps.comp \<open>Cod (SPN \<mu>)\<close>
proof
interpret s: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>cod \<mu>\<close>
using assms by unfold_locales auto
show "Maps.span (Leg0 (Cod (SPN \<mu>))) (Leg1 (Cod (SPN \<mu>)))"
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
\<open>\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (cod \<mu>)\<rbrakk>\<rbrakk>\<rparr>\<close>
using assms Cod.span_in_category_axioms SPN_def by simp
show 1: "Span.arr (SPN \<mu>)"
proof (unfold Span.arr_char)
show "arrow_of_spans Maps.comp (SPN \<mu>)"
proof (unfold_locales)
show "Maps.in_hom (Chn (SPN \<mu>)) 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 \<odot> Chn (SPN \<mu>) = Dom.leg0"
unfolding SPN_def
using assms spn_props [of \<mu>] Maps.comp_CLS [of "tab\<^sub>0 (cod \<mu>)" "spn \<mu>"] by simp
show "Cod.leg1 \<odot> Chn (SPN \<mu>) = Dom.leg1"
unfolding SPN_def
using assms spn_props [of \<mu>] Maps.comp_CLS [of "tab\<^sub>1 (cod \<mu>)" "spn \<mu>"] by simp
qed
qed
show "Span.dom (SPN \<mu>) = SPN\<^sub>0 (dom \<mu>)"
using assms 1 Span.dom_char Dom'.apex_def SPN_def by simp
show "Span.cod (SPN \<mu>) = SPN\<^sub>0 (cod \<mu>)"
using assms 1 Span.cod_char Cod'.apex_def SPN_def by simp
qed
interpretation SPN: "functor" V Span.vcomp SPN
proof
show "\<And>\<mu>. \<not> arr \<mu> \<Longrightarrow> SPN \<mu> = Span.null"
unfolding SPN_def by simp
show 1: "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.arr (SPN \<mu>)"
using SPN_in_hom by auto
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.dom (SPN \<mu>) = SPN (dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
have 1: "Maps.arr (Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>)"
proof -
have "src (tab\<^sub>0 (dom \<mu>)) \<in> Collect obj"
using \<mu> by simp
moreover have "src \<mu> \<in> Collect obj"
using \<mu> by simp
moreover have "\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (local.dom \<mu>))) (src \<mu>)"
proof -
have "\<guillemotleft>tab\<^sub>0 (dom \<mu>) : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src \<mu>\<guillemotright>"
using \<mu> by simp
moreover have "is_left_adjoint (tab\<^sub>0 (dom \<mu>))"
using \<mu> tab\<^sub>0_simps [of "dom \<mu>"] by auto
ultimately show ?thesis by auto
qed
ultimately show ?thesis by simp
qed
have "\<lbrakk>spn (dom \<mu>)\<rbrakk> = \<lbrakk>src (tab\<^sub>0 (dom \<mu>))\<rbrakk>"
using \<mu> spn_ide iso_class_eqI by auto
hence "SPN (dom \<mu>) = SPN\<^sub>0 (dom \<mu>)"
unfolding SPN_def
using \<mu> 1 Maps.dom_char by simp
thus "Span.dom (SPN \<mu>) = SPN (dom \<mu>)"
using \<mu> SPN_in_hom by auto
qed
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.cod (SPN \<mu>) = SPN (cod \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
have 1: "Maps.arr (Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>)"
proof -
have "src (tab\<^sub>0 (cod \<mu>)) \<in> Collect obj"
using \<mu> by simp
moreover have "src \<mu> \<in> Collect obj"
using \<mu> by simp
moreover have "\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (cod \<mu>))) (src \<mu>)"
proof -
have "\<guillemotleft>tab\<^sub>0 (cod \<mu>) : src (tab\<^sub>0 (cod \<mu>)) \<rightarrow> src \<mu>\<guillemotright>"
using \<mu> by simp
moreover have "is_left_adjoint (tab\<^sub>0 (cod \<mu>))"
using \<mu> by simp
ultimately show ?thesis by auto
qed
ultimately show ?thesis by simp
qed
have "\<lbrakk>spn (cod \<mu>)\<rbrakk> = \<lbrakk>src (tab\<^sub>0 (cod \<mu>))\<rbrakk>"
using \<mu> spn_ide iso_class_eqI by auto
hence "SPN (cod \<mu>) = SPN\<^sub>0 (cod \<mu>)"
unfolding SPN_def
using \<mu> 1 Maps.dom_char by simp
thus "Span.cod (SPN \<mu>) = SPN (cod \<mu>)"
using \<mu> SPN_in_hom by auto
qed
show "\<And>\<nu> \<mu>. seq \<nu> \<mu> \<Longrightarrow> SPN (\<nu> \<cdot> \<mu>) = SPN \<nu> \<bullet> SPN \<mu>"
proof -
fix \<mu> \<nu>
assume seq: "seq \<nu> \<mu>"
have "Dom (SPN (\<nu> \<cdot> \<mu>)) = Dom (SPN \<nu> \<bullet> SPN \<mu>)"
using seq 1 Span.vcomp_def Span.arr_char
by (elim seqE, simp)
moreover have "Cod (SPN (\<nu> \<cdot> \<mu>)) = Cod (SPN \<nu> \<bullet> SPN \<mu>)"
using seq 1 Span.vcomp_def Span.arr_char
by (elim seqE, simp)
moreover have "Chn (SPN (\<nu> \<cdot> \<mu>)) = Chn (SPN \<nu> \<bullet> SPN \<mu>)"
proof -
have *: "\<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk> = Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
proof
show "\<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk> \<subseteq> Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
proof
fix h
assume h: "h \<in> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
show "h \<in> Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
proof -
have 1: "spn \<nu> \<in> \<lbrakk>spn \<nu>\<rbrakk>"
using seq ide_in_iso_class by auto
moreover have 2: "spn \<mu> \<in> \<lbrakk>spn \<mu>\<rbrakk>"
using seq ide_in_iso_class by auto
moreover have "spn \<nu> \<star> spn \<mu> \<cong> h"
proof -
have "spn \<nu> \<star> spn \<mu> \<cong> spn (\<nu> \<cdot> \<mu>)"
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 \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk> \<subseteq> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
proof
fix h
assume h: "h \<in> Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
show "h \<in> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
proof -
obtain f g where 1: "g \<in> \<lbrakk>spn \<nu>\<rbrakk> \<and> f \<in> \<lbrakk>spn \<mu>\<rbrakk> \<and> g \<star> f \<cong> h"
using h Maps.Comp_def [of "iso_class (spn \<nu>)" "iso_class (spn \<mu>)"]
iso_class_def iso_class_elems_isomorphic
by blast
have fg: "g \<cong> spn \<nu> \<and> f \<cong> spn \<mu> \<and> g \<star> f \<cong> h"
proof -
have "spn \<nu> \<in> \<lbrakk>spn \<nu>\<rbrakk> \<and> spn \<mu> \<in> \<lbrakk>spn \<mu>\<rbrakk>"
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 \<star> f \<in> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
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 \<nu> \<bullet> SPN \<mu>) =
Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk> \<odot>
Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>))) \<lbrakk>spn \<mu>\<rbrakk>"
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 \<mu>))) (src (tab\<^sub>0 (cod \<nu>)))
(Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>)"
proof -
have "Maps.seq (Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>))) \<lbrakk>spn \<mu>\<rbrakk>)"
proof
show "Maps.in_hom (Maps.MkArr (src (tab\<^sub>0 (local.dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>)))
\<lbrakk>spn \<mu>\<rbrakk>)
(Maps.MkIde (src (tab\<^sub>0 (dom \<mu>))))
(Maps.MkIde (src (tab\<^sub>0 (cod \<mu>))))"
proof -
have "src (tab\<^sub>0 (dom \<mu>)) \<in> Collect obj"
using in_hhom_def seq by auto
moreover have "src (tab\<^sub>0 (cod \<mu>)) \<in> Collect obj"
using seq by auto
moreover have "\<lbrakk>spn \<mu>\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>)))"
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 \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk>)
(Maps.MkIde (src (tab\<^sub>0 (cod \<mu>))))
(Maps.MkIde (src (tab\<^sub>0 (cod \<nu>))))"
proof -
have "src (tab\<^sub>0 (cod \<mu>)) \<in> Collect obj"
using in_hhom_def seq by auto
moreover have "src (tab\<^sub>0 (cod \<nu>)) \<in> Collect obj"
using seq by auto
moreover have "\<lbrakk>spn \<nu>\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>)))"
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 \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk>"
"Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>))) \<lbrakk>spn \<mu>\<rbrakk>"]
by simp
qed
also have "... = Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
using * by simp
also have "... = Chn (SPN (\<nu> \<cdot> \<mu>))"
using seq SPN_def Span.vcomp_def
by (elim seqE, simp)
finally show ?thesis by simp
qed
ultimately show "SPN (\<nu> \<cdot> \<mu>) = SPN \<nu> \<bullet> SPN \<mu>" 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 "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.isomorphic (SPN (src \<mu>)) (Span.src (SPN \<mu>))"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
let ?src = "Maps.MkIde (src \<mu>)"
have src: "Maps.ide ?src"
using \<mu> by simp
interpret src: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>src \<mu>\<close>
using \<mu> by unfold_locales auto
interpret src: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>src \<mu>\<close> src.tab \<open>tab\<^sub>0 (src \<mu>)\<close> \<open>tab\<^sub>1 (src \<mu>)\<close>
\<open>src \<mu>\<close> src.tab \<open>tab\<^sub>0 (src \<mu>)\<close> \<open>tab\<^sub>1 (src \<mu>)\<close>
\<open>src \<mu>\<close>
using src.is_arrow_of_tabulations_in_maps by simp
interpret src: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?src, Leg1 = ?src\<rparr>\<close>
using src by (unfold_locales, simp)
let ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (src \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
have tab\<^sub>0_src: "\<guillemotleft>tab\<^sub>0 (src \<mu>) : src (tab\<^sub>0 (src \<mu>)) \<rightarrow> src \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>0 (src \<mu>)) \<and> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>0: "Maps.arr ?tab\<^sub>0"
using \<mu> Maps.arr_MkArr tab\<^sub>0_src by blast
let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (src \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk>"
have tab\<^sub>1_src: "\<guillemotleft>tab\<^sub>1 (src \<mu>) : src (tab\<^sub>0 (src \<mu>)) \<rightarrow> src \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>1 (src \<mu>)) \<and> \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>1: "Maps.arr ?tab\<^sub>1"
using \<mu> Maps.arr_MkArr tab\<^sub>1_src by blast
interpret tab: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<close>
using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp)
have "src \<mu> \<star> tab\<^sub>0 (src \<mu>) \<cong> tab\<^sub>0 (src \<mu>)"
using \<mu> iso_lunit isomorphic_def
by (metis lunit_in_hom(2) src.ide_u src.u_simps(3) src_src)
hence "src \<mu> \<star> tab\<^sub>0 (src \<mu>) \<cong> tab\<^sub>1 (src \<mu>)"
using \<mu> src.obj_has_symmetric_tab isomorphic_transitive by blast
have "\<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (src \<mu>))) (src \<mu>)"
using \<mu> tab\<^sub>0_src by blast
have "\<lbrakk>src \<mu>\<rbrakk> \<in> Maps.Hom (src \<mu>) (src \<mu>)"
proof -
have "\<guillemotleft>src \<mu> : src \<mu> \<rightarrow> src \<mu>\<guillemotright> \<and> is_left_adjoint (src \<mu>) \<and> \<lbrakk>src \<mu>\<rbrakk> = \<lbrakk>src \<mu>\<rbrakk>"
using \<mu> obj_is_self_adjoint by simp
thus ?thesis by auto
qed
interpret SPN_src: arrow_of_spans Maps.comp \<open>SPN (src \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have SPN_src: "SPN (src \<mu>) =
\<lparr>Chn = Maps.MkArr (src (tab\<^sub>0 (src \<mu>))) (src (tab\<^sub>0 (src \<mu>))) \<lbrakk>spn (src \<mu>)\<rbrakk>,
Dom = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<rparr>"
unfolding SPN_def using \<mu> by simp
interpret src_SPN: arrow_of_spans Maps.comp \<open>Span.src (SPN \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have src_SPN: "Span.src (SPN \<mu>) =
\<lparr>Chn = ?src,
Dom = \<lparr>Leg0 = ?src, Leg1 = ?src\<rparr>,
Cod = \<lparr>Leg0 = ?src, Leg1 = ?src\<rparr>\<rparr>"
proof -
let ?tab\<^sub>0_dom = "Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>"
have "Maps.arr ?tab\<^sub>0_dom"
proof -
have "\<guillemotleft>tab\<^sub>0 (dom \<mu>) : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>0 (dom \<mu>)) \<and> \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>"
using \<mu> by simp
thus ?thesis
using \<mu> Maps.arr_MkArr by blast
qed
thus ?thesis
using \<mu> Maps.cod_char Span.src_def by simp
qed
text \<open>
The idea of the proof is that @{term "iso_class (tab\<^sub>0 (src \<mu>))"} is invertible
in \<open>Maps(B)\<close> and determines an invertible arrow of spans from @{term "SPN (src \<mu>)"}
to @{term "Span.src (SPN \<mu>)"}.
\<close>
let ?\<phi> = "\<lparr>Chn = ?tab\<^sub>0, Dom = Dom (SPN (src \<mu>)), Cod = Cod (Span.src (SPN \<mu>))\<rparr>"
interpret \<phi>: arrow_of_spans Maps.comp ?\<phi>
apply (unfold_locales, simp_all)
proof -
show "Maps.in_hom ?tab\<^sub>0 SPN_src.dom.apex src_SPN.cod.apex"
using \<mu> 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 \<odot> ?tab\<^sub>0 = SPN_src.dom.leg0"
proof -
have "Maps.seq src_SPN.cod.leg0 ?tab\<^sub>0"
using \<mu> src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>0 (src \<mu>) \<in> Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "src \<mu>"]
ide_in_iso_class [of "tab\<^sub>0 (src \<mu>)"] \<open>src \<mu> \<star> tab\<^sub>0 (src \<mu>) \<cong> tab\<^sub>0 (src \<mu>)\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>\<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (src \<mu>))) (src \<mu>)\<close>
\<open>\<lbrakk>src \<mu>\<rbrakk> \<in> Maps.Hom (src \<mu>) (src \<mu>)\<close>
by meson
qed
ultimately show ?thesis
using \<mu> Maps.comp_char [of src_SPN.cod.leg0 ?tab\<^sub>0] src_SPN by simp
qed
show "src_SPN.cod.leg1 \<odot> ?tab\<^sub>0 = SPN_src.dom.leg1"
proof -
have "Maps.seq src_SPN.cod.leg1 ?tab\<^sub>0"
using \<mu> src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>1 (src \<mu>) \<in> Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "src \<mu>"]
ide_in_iso_class [of "tab\<^sub>0 (src \<mu>)"]
\<open>isomorphic (src \<mu> \<star> tab\<^sub>0 (src \<mu>)) (tab\<^sub>1 (src \<mu>))\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>\<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (src \<mu>))) (src \<mu>)\<close>
\<open>\<lbrakk>src \<mu>\<rbrakk> \<in> Maps.Hom (src \<mu>) (src \<mu>)\<close>
by meson
qed
ultimately show ?thesis
using \<mu> Maps.comp_char [of src_SPN.cod.leg1 ?tab\<^sub>0] src_SPN by simp
qed
qed
have "Span.in_hom ?\<phi> (SPN (src \<mu>)) (Span.src (SPN \<mu>))"
using \<mu> tab\<^sub>0 spn_ide [of "src \<mu>"] iso_class_eqI
Span.arr_char Span.dom_char Span.cod_char \<phi>.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 \<mu> 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 \<mu>)) (Span.src (SPN \<mu>))"
using Span.isomorphic_def Span.iso_char by auto
qed
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.isomorphic (SPN (trg \<mu>)) (Span.trg (SPN \<mu>))"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
let ?trg = "Maps.MkIde (trg \<mu>)"
have trg: "Maps.ide ?trg"
using \<mu> by simp
interpret trg: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>trg \<mu>\<close>
using \<mu> by unfold_locales auto
interpret trg: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?trg, Leg1 = ?trg\<rparr>\<close>
using trg by (unfold_locales, simp)
let ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>) \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
have tab\<^sub>0_trg: "\<guillemotleft>tab\<^sub>0 (trg \<mu>) : src (tab\<^sub>0 (trg \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>0 (trg \<mu>)) \<and> \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>0: "Maps.arr ?tab\<^sub>0"
using \<mu> Maps.arr_MkArr tab\<^sub>0_trg by blast
let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>) \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
have tab\<^sub>1_trg: "\<guillemotleft>tab\<^sub>1 (trg \<mu>) : src (tab\<^sub>0 (trg \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>1 (trg \<mu>)) \<and> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>1: "Maps.arr ?tab\<^sub>1"
using \<mu> Maps.arr_MkArr tab\<^sub>1_trg by blast
interpret tab: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<close>
using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp)
have "trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>0 (trg \<mu>)"
proof -
have "\<guillemotleft>\<l>[tab\<^sub>1 (trg \<mu>)] : trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<Rightarrow> tab\<^sub>1 (trg \<mu>)\<guillemotright>"
using \<mu> by simp
moreover have "iso \<l>[tab\<^sub>1 (trg \<mu>)]"
using \<mu> iso_lunit by simp
ultimately have "trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)"
using isomorphic_def by auto
also have "tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>0 (trg \<mu>)"
using \<mu> trg.obj_has_symmetric_tab isomorphic_symmetric by auto
finally show ?thesis by blast
qed
hence "trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)"
using \<mu> trg.obj_has_symmetric_tab isomorphic_transitive by blast
have "\<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>)"
proof -
have "\<guillemotleft>tab\<^sub>1 (trg \<mu>) : src (tab\<^sub>0 (trg \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and> is_left_adjoint (tab\<^sub>0 (trg \<mu>)) \<and>
\<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
using \<mu> by simp
thus ?thesis by auto
qed
have "\<lbrakk>trg \<mu>\<rbrakk> \<in> Maps.Hom (trg \<mu>) (trg \<mu>)"
proof -
have "\<guillemotleft>trg \<mu> : trg \<mu> \<rightarrow> trg \<mu>\<guillemotright> \<and> is_left_adjoint (trg \<mu>) \<and> \<lbrakk>trg \<mu>\<rbrakk> = \<lbrakk>trg \<mu>\<rbrakk>"
using \<mu> obj_is_self_adjoint by simp
thus ?thesis by auto
qed
interpret SPN_trg: arrow_of_spans Maps.comp \<open>SPN (trg \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have SPN_trg: "SPN (trg \<mu>) =
\<lparr>Chn = Maps.MkArr (src (tab\<^sub>1 (trg \<mu>))) (src (tab\<^sub>1 (trg \<mu>))) \<lbrakk>spn (trg \<mu>)\<rbrakk>,
Dom = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<rparr>"
unfolding SPN_def using \<mu> by simp
interpret trg_SPN: arrow_of_spans Maps.comp \<open>Span.trg (SPN \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have trg_SPN: "Span.trg (SPN \<mu>) = \<lparr>Chn = ?trg,
Dom = \<lparr>Leg0 = ?trg, Leg1 = ?trg\<rparr>,
Cod = \<lparr>Leg0 = ?trg, Leg1 = ?trg\<rparr>\<rparr>"
proof -
let ?tab\<^sub>1_dom = "Maps.MkArr (src (tab\<^sub>1 (dom \<mu>))) (trg \<mu>) \<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>"
have "Maps.arr ?tab\<^sub>1_dom"
proof -
have "\<guillemotleft>tab\<^sub>1 (dom \<mu>) : src (tab\<^sub>1 (dom \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>1 (dom \<mu>)) \<and> \<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>"
using \<mu> by simp
thus ?thesis
using \<mu> Maps.arr_MkArr by blast
qed
thus ?thesis
using \<mu> Maps.cod_char Span.trg_def by simp
qed
let ?\<phi> = "\<lparr>Chn = ?tab\<^sub>1, Dom = Dom (SPN (trg \<mu>)), Cod = Cod (Span.trg (SPN \<mu>))\<rparr>"
interpret \<phi>: arrow_of_spans Maps.comp ?\<phi>
apply (unfold_locales, simp_all)
proof -
show "Maps.in_hom ?tab\<^sub>1 SPN_trg.dom.apex trg_SPN.cod.apex"
using \<mu> 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 \<mu> trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>1 (trg \<mu>) \<in> Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "trg \<mu>"]
ide_in_iso_class [of "tab\<^sub>1 (trg \<mu>)"] \<open>trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>iso_class (tab\<^sub>1 (trg \<mu>)) \<in> Maps.Hom (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>)\<close>
\<open>iso_class (trg \<mu>) \<in> Maps.Hom (trg \<mu>) (trg \<mu>)\<close>
by meson
qed
moreover have "\<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
using \<mu> iso_class_eqI trg.obj_has_symmetric_tab by auto
ultimately show ?thesis
using \<mu> Maps.comp_char [of trg_SPN.cod.leg0 ?tab\<^sub>1] trg_SPN
by simp
qed
show "trg_SPN.cod.leg1 \<odot> ?tab\<^sub>1 = SPN_trg.dom.leg1"
proof -
have "Maps.seq trg_SPN.cod.leg1 ?tab\<^sub>1"
using \<mu> trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>1 (trg \<mu>) \<in> Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "trg \<mu>"]
ide_in_iso_class [of "tab\<^sub>1 (trg \<mu>)"] \<open>trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>\<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>)\<close>
\<open>\<lbrakk>trg \<mu>\<rbrakk> \<in> Maps.Hom (trg \<mu>) (trg \<mu>)\<close>
by meson
qed
ultimately show ?thesis
using \<mu> Maps.comp_char [of trg_SPN.cod.leg1 ?tab\<^sub>1] trg_SPN by simp
qed
qed
have \<phi>: "Span.in_hom ?\<phi> (SPN (trg \<mu>)) (Span.trg (SPN \<mu>))"
using \<mu> tab\<^sub>0 spn_ide [of "trg \<mu>"] iso_class_eqI
Span.arr_char Span.dom_char Span.cod_char \<phi>.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 \<mu> 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 "\<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> 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 \<mu>)) (Span.trg (SPN \<mu>))"
using \<phi> 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 \<open>
To complete the proof that \<open>SPN\<close> is a pseudofunctor, we need to obtain a natural
isomorphism \<open>\<Phi>\<close>, whose component at \<open>(r, s)\<close> is an isomorphism \<open>\<Phi> (r, s)\<close>
from the horizontal composite \<open>SPN r \<circ> SPN s\<close> to \<open>SPN (r \<star> s)\<close> in \<open>Span(Maps(B))\<close>,
and we need to prove that the coherence conditions are satisfied.
We have shown that the tabulations of \<open>r\<close> and \<open>s\<close> compose to yield a tabulation of \<open>r \<star> s\<close>.
Since tabulations of the same arrow are equivalent, this tabulation must be equivalent
to the chosen tabulation of \<open>r \<star> s\<close>. We therefore obtain an equivalence map from the
apex of \<open>SPN r \<circ> SPN s\<close> to the apex of \<open>SPN (r \<star> s)\<close> which commutes with the
legs of these spans up to isomorphism. This equivalence map determines an invertible
arrow in \<open>Span(Maps(B))\<close>. Moreover, by property \<open>T2\<close>, any two such equivalence maps are
connected by a unique 2-cell, which is consequently an isomorphism. This shows that
the arrow in \<open>Span(Maps(B))\<close> is uniquely determined, which fact we can exploit to establish
the required coherence conditions.
\<close>
locale two_composable_identities_in_bicategory_of_spans =
bicategory_of_spans V H \<a> \<i> src trg +
r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r +
s: identity_in_bicategory_of_spans V H \<a> \<i> src trg s
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and s :: 'a +
assumes composable: "src r = trg s"
begin
notation isomorphic (infix "\<cong>" 50)
interpretation r: arrow_in_bicategory_of_spans V H \<a> \<i> src trg r r r
by unfold_locales auto
interpretation r: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
r
using r.is_arrow_of_tabulations_in_maps by simp
interpretation s: arrow_in_bicategory_of_spans V H \<a> \<i> src trg s s s
by unfold_locales auto
interpretation s: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
s
using s.is_arrow_of_tabulations_in_maps by simp
sublocale identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>r \<star> s\<close>
apply unfold_locales by (simp add: composable)
sublocale horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
r s
using composable by unfold_locales auto
abbreviation p\<^sub>0 where "p\<^sub>0 \<equiv> \<rho>\<sigma>.p\<^sub>0"
abbreviation p\<^sub>1 where "p\<^sub>1 \<equiv> \<rho>\<sigma>.p\<^sub>1"
text \<open>
We will take as the composition isomorphism from \<open>SPN r \<circ> SPN s\<close> to \<open>SPN (r \<star> s)\<close>
the arrow of tabulations, induced by the identity \<open>r \<star> s\<close>, from the composite of
the chosen tabulations of \<open>r\<close> and \<open>s\<close> to the chosen tabulation of \<open>r \<star> s\<close>.
As this arrow of tabulations is induced by an identity, it is an equivalence map.
\<close>
interpretation cmp: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> \<rho>\<sigma>.tab \<open>tab\<^sub>0 s \<star> \<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 r \<star> \<rho>\<sigma>.p\<^sub>1\<close>
\<open>r \<star> s\<close> tab \<open>tab\<^sub>0 (r \<star> s)\<close> \<open>tab\<^sub>1 (r \<star> s)\<close>
\<open>r \<star> s\<close>
using composable
by unfold_locales auto
lemma cmp_interpretation:
shows "identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
(r \<star> s) \<rho>\<sigma>.tab (tab\<^sub>0 s \<star> \<rho>\<sigma>.p\<^sub>0) (tab\<^sub>1 r \<star> \<rho>\<sigma>.p\<^sub>1)
(r \<star> s) tab (tab\<^sub>0 (r \<star> s)) (tab\<^sub>1 (r \<star> s))
(r \<star> s)"
..
definition cmp
where "cmp = cmp.chine"
lemma cmp_props:
shows "\<guillemotleft>cmp : src \<rho>\<sigma>.tab \<rightarrow> src tab\<guillemotright>"
and "\<guillemotleft>cmp : cmp \<Rightarrow> cmp\<guillemotright>"
and "equivalence_map cmp"
and "tab\<^sub>0 (r \<star> s) \<star> cmp \<cong> tab\<^sub>0 s \<star> \<rho>\<sigma>.p\<^sub>0"
and "tab\<^sub>1 (r \<star> s) \<star> cmp \<cong> tab\<^sub>1 r \<star> \<rho>\<sigma>.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 "\<guillemotleft>cmp : src \<rho>\<sigma>.tab \<rightarrow> src tab\<guillemotright>"
and "\<guillemotleft>cmp : cmp \<Rightarrow> cmp\<guillemotright>"
using cmp_props by auto
lemma cmp_simps [simp]:
shows "arr cmp" and "ide cmp"
and "src cmp = src \<rho>\<sigma>.tab" and "trg cmp = src tab"
and "dom cmp = cmp" and "cod cmp = cmp"
using cmp_props equivalence_map_is_ide by auto
text \<open>
Now we have to use the above properties of the underlying bicategory to
exhibit the composition isomorphisms as actual arrows in \<open>Span(Maps(B))\<close>
and to prove the required naturality and coherence conditions.
\<close>
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 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 \<open>SPN r \<circ> SPN s\<close>
using composable Span.ide_char [of "SPN r \<circ> SPN s"] by simp
interpretation SPN_r_SPN_s: identity_arrow_of_spans Maps.comp \<open>SPN r \<circ> SPN s\<close>
using composable Span.ide_char [of "SPN r \<circ> SPN s"]
by (unfold_locales, simp)
interpretation SPN_rs: arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using composable Span.arr_char r.base_simps(2) s.base_simps(2) by blast
interpretation SPN_rs: identity_arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using composable Span.ide_char SPN.preserves_ide r.is_ide s.is_ide
by (unfold_locales, simp)
text \<open>
The following are the legs (as arrows of \<open>Maps\<close>) of the spans \<open>SPN r\<close> and \<open>SPN s\<close>.
\<close>
definition R\<^sub>0 where "R\<^sub>0 = \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk>"
definition R\<^sub>1 where "R\<^sub>1 = \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>"
definition S\<^sub>0 where "S\<^sub>0 = \<lbrakk>\<lbrakk>tab\<^sub>0 s\<rbrakk>\<rbrakk>"
definition S\<^sub>1 where "S\<^sub>1 = \<lbrakk>\<lbrakk>tab\<^sub>1 s\<rbrakk>\<rbrakk>"
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 \<open>
The apex of the composite span @{term "SPN r \<circ> SPN s"} (defined in terms of pullback)
coincides with the apex of the composite tabulation \<open>\<rho>\<sigma>\<close> (defined using
the chosen tabulation of \<open>(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s)\<close>). We need this to be true in order
to define the compositor of a pseudofunctor from the underlying bicategory \<open>B\<close>
to \<open>Span(Maps(B))\<close>. It is only true if we have carefully chosen pullbacks in \<open>Maps(B)\<close>
in order to ensure the relationship with the chosen tabulations.
\<close>
lemma SPN_r_SPN_s_apex_eq:
shows "SPN_r_SPN_s.apex = Maps.MkIde (src \<rho>\<sigma>.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 \<noteq> Maps.Null"
using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
moreover have "Maps.Dom SPN_r_SPN_s.leg0 = src \<rho>\<sigma>.tab"
proof -
interpret REP_S\<^sub>1: map_in_bicategory V H \<a> \<i> src trg \<open>Maps.REP S\<^sub>1\<close>
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 \<a> \<i> src trg \<open>Maps.REP R\<^sub>0\<close>
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 \<odot> 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 \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1))"
proof -
have "ide ((Maps.REP R\<^sub>0)\<^sup>* \<star> 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 \<odot> \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>) =
src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1))"
proof -
have "Maps.arr (\<lbrakk>\<lbrakk>prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\<rbrakk>\<rbrakk>)"
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 \<lbrakk>\<lbrakk>prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\<rbrakk>\<rbrakk>"
proof -
have "Maps.Cod \<lbrakk>\<lbrakk>prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\<rbrakk>\<rbrakk> =
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 \<odot> \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>) =
Maps.Dom \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>"
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>* \<star> 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>* \<star> Maps.REP S\<^sub>1))"
by simp
also have "... = src (tab\<^sub>0 (r.s\<^sub>0\<^sup>* \<star> s.s\<^sub>1))"
proof -
interpret r\<^sub>0's\<^sub>1: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>r.s\<^sub>0\<^sup>* \<star> s.s\<^sub>1\<close>
using composable by (unfold_locales, simp)
have "(Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1 \<cong> r.s\<^sub>0\<^sup>* \<star> s.s\<^sub>1"
proof -
have "(Maps.REP R\<^sub>0)\<^sup>* \<cong> 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 \<cong> 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 \<cong> s.s\<^sub>1"
unfolding S\<^sub>1_def
using Maps.REP_CLS s.leg1_is_map composable by force
moreover have "\<exists>a. a \<cong> (tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s \<and> (Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1 \<cong> 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 \<rho>\<sigma>.tab"
- using VV.ide_char VV.arr_char composable Span.hcomp_def \<rho>\<sigma>.tab_simps(2) by auto
+ using VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C composable Span.hcomp_def \<rho>\<sigma>.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 \<open>
We will be taking the arrow @{term "CLS cmp"} of \<open>Maps\<close> as the composition isomorphism from
@{term "SPN r \<circ> SPN s"} to @{term "SPN (r \<star> s)"}. The following result shows that it
has the right domain and codomain for that purpose.
\<close>
lemma iso_class_cmp_in_hom:
shows "Maps.in_hom (Maps.MkArr (src \<rho>\<sigma>.tab) (src tab) \<lbrakk>cmp\<rbrakk>)
SPN_r_SPN_s.apex SPN_rs.apex"
and "Maps.in_hom \<lbrakk>\<lbrakk>cmp\<rbrakk>\<rbrakk> SPN_r_SPN_s.apex SPN_rs.apex"
proof -
show "Maps.in_hom (Maps.MkArr (src \<rho>\<sigma>.tab) (src tab) \<lbrakk>cmp\<rbrakk>)
SPN_r_SPN_s.apex SPN_rs.apex"
proof -
have "obj (src \<rho>\<sigma>.tab)"
using obj_src \<rho>\<sigma>.tab_in_hom by blast
moreover have "obj (src tab)"
using obj_src by simp
moreover have "\<lbrakk>cmp\<rbrakk> \<in> Maps.Hom (src \<rho>\<sigma>.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 \<rho>\<sigma>.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 \<lbrakk>\<lbrakk>cmp\<rbrakk>\<rbrakk> 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 \<a> \<i> src trg \<open>(Maps.REP R\<^sub>0)\<^sup>*\<close> \<open>Maps.REP S\<^sub>1\<close>
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 \<a> \<i> src trg \<open>(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s\<close>
by (unfold_locales, simp add: composable)
lemma prj_tab_agreement:
shows "(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s \<cong> (Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1"
and "\<rho>\<sigma>.p\<^sub>0 \<cong> prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)"
and "\<rho>\<sigma>.p\<^sub>1 \<cong> prj\<^sub>1 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)"
proof -
show 1: "(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s \<cong> (Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1"
proof -
have "(tab\<^sub>0 r)\<^sup>* \<cong> (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 \<cong> 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 "\<rho>\<sigma>.p\<^sub>0 \<cong> tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)"
using 1 R\<^sub>0'S\<^sub>1.isomorphic_implies_same_tab isomorphic_reflexive by auto
show "\<rho>\<sigma>.p\<^sub>1 \<cong> tab\<^sub>1 ((Maps.REP R\<^sub>0)\<^sup>* \<star> 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 \<rho>\<sigma>.p\<^sub>0)"
proof -
have "Span.chine_hcomp (SPN r) (SPN s) =
Maps.MkIde (src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> 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 \<rho>\<sigma>.p\<^sub>0)"
using prj_tab_agreement isomorphic_implies_hpar(3) by force
finally show ?thesis by simp
qed
end
text \<open>
The development above focused on two specific composable 1-cells in bicategory \<open>B\<close>.
Here we reformulate those results as statements about the entire bicategory.
\<close>
context bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 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 \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<circ> snd \<mu>\<nu>\<close>
..
interpretation SPNoH: composite_functor VV.comp V
Span.vcomp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> SPN
..
text \<open>
Given arbitrary composable 1-cells \<open>r\<close> and \<open>s\<close>, obtain an arrow of spans in \<open>Maps\<close>
having the isomorphism class of \<open>rs.cmp\<close> as its chine.
\<close>
definition CMP
where "CMP r s \<equiv>
\<lparr>Chn = \<lbrakk>\<lbrakk>two_composable_identities_in_bicategory_of_spans.cmp V H \<a> \<i> src trg r s\<rbrakk>\<rbrakk>,
Dom = Dom (SPN r \<circ> SPN s), Cod = Dom (SPN (r \<star> s))\<rparr>"
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
+ using assms VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
interpret rs: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg r s
- using rs VV.ide_char VV.arr_char by unfold_locales auto
+ using rs VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by unfold_locales auto
interpret cmp: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> rs.\<rho>\<sigma>.tab \<open>tab\<^sub>0 s \<star> rs.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 r \<star> rs.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>r \<star> s\<close> rs.tab \<open>tab\<^sub>0 (r \<star> s)\<close> \<open>tab\<^sub>1 (r \<star> s)\<close>
\<open>r \<star> s\<close>
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 \<open>SPN r \<circ> SPN s\<close>
using rs.composable Span.ide_char [of "SPN r \<circ> SPN s"] by simp
interpret SPN_r_SPN_s: identity_arrow_of_spans Maps.comp \<open>SPN r \<circ> SPN s\<close>
using rs.composable Span.ide_char [of "SPN r \<circ> SPN s"]
by (unfold_locales, simp)
interpret SPN_rs: arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using Span.arr_char rs.is_ide SPN.preserves_arr by blast
interpret SPN_rs: identity_arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using Span.ide_char rs.is_ide SPN.preserves_ide
by (unfold_locales, simp)
interpret Dom: span_in_category Maps.comp \<open>Dom (CMP r s)\<close>
by (unfold_locales, simp add: CMP_def)
interpret Cod: span_in_category Maps.comp \<open>Cod (CMP r s)\<close>
proof -
(* TODO: I don't understand what makes this so difficult. *)
have "\<guillemotleft>tab\<^sub>0 (r \<star> s) : src (tab\<^sub>0 (r \<star> s)) \<rightarrow> src s\<guillemotright> \<and> is_left_adjoint (tab\<^sub>0 (r \<star> s)) \<and>
\<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> = \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk>"
by simp
hence "\<exists>f. \<guillemotleft>f : src (tab\<^sub>0 (r \<star> s)) \<rightarrow> src s\<guillemotright> \<and> is_left_adjoint f \<and> \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> = \<lbrakk>f\<rbrakk>"
by blast
moreover have "\<exists>f. \<guillemotleft>f : src (tab\<^sub>0 (r \<star> s)) \<rightarrow> trg r\<guillemotright> \<and> is_left_adjoint f \<and>
\<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk> = \<lbrakk>f\<rbrakk>"
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 \<a> \<i> src trg \<open>(Maps.REP rs.R\<^sub>0)\<^sup>*\<close> \<open>Maps.REP rs.S\<^sub>1\<close>
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 \<a> \<i> src trg \<open>(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s\<close>
by (unfold_locales, simp add: rs.composable)
text \<open>
Here we obtain explicit formulas for the legs and apex of \<open>SPN_r_SPN_s\<close> and \<open>SPN_rs\<close>.
\<close>
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.\<rho>\<sigma>.tab)"
using rs.SPN_r_SPN_s_apex_eq by auto
have SPN_rs_leg0_eq: "SPN_rs.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk>\<rbrakk>"
unfolding SPN_def using rs by simp
have SPN_rs_leg1_eq: "SPN_rs.leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk>\<rbrakk>"
unfolding SPN_def using rs by simp
have "SPN_rs.apex = Maps.MkIde (src (tab_of_ide (r \<star> s)))"
using SPN_rs.dom.apex_def Maps.dom_char SPN_rs_leg0_eq SPN_rs.dom.leg_simps(1)
by simp
text \<open>
The composition isomorphism @{term "CMP r s"} is an arrow of spans in \<open>Maps(B)\<close>.
\<close>
interpret arrow_of_spans Maps.comp \<open>CMP r s\<close>
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 \<odot> 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 \<odot> 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 \<odot> 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 \<odot> Chn (CMP r s)) = Maps.Map Dom.leg0"
proof -
have "Maps.Map (Cod.leg0 \<odot> Chn (CMP r s)) = Maps.Comp \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk>"
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 \<lbrakk>tab\<^sub>0 s\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>0\<rbrakk>"
proof -
have "Maps.Comp \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk> = \<lbrakk>tab\<^sub>0 (r \<star> s) \<star> rs.cmp\<rbrakk>"
using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def
by auto
also have "... = Maps.Comp \<lbrakk>tab\<^sub>0 s\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>0\<rbrakk>"
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 "\<lbrakk>prj\<^sub>0 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\<rbrakk> = \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>0\<rbrakk>"
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.\<rho>\<sigma>.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 \<odot> 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 \<odot> 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 \<odot> 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 \<odot> Chn (CMP r s)) = Maps.Map Dom.leg1"
proof -
have "Maps.Map (Cod.leg1 \<odot> Chn (CMP r s)) = Maps.Comp \<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk>"
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 \<lbrakk>tab\<^sub>1 r\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>1\<rbrakk>"
proof -
have "Maps.Comp \<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk> = \<lbrakk>tab\<^sub>1 (r \<star> s) \<star> rs.cmp\<rbrakk>"
using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def
by auto
also have "... = Maps.Comp \<lbrakk>tab\<^sub>1 r\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using Maps.Comp_eq_iso_class_memb
Maps.CLS_hcomp [of "tab\<^sub>1 r" rs.\<rho>\<sigma>.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 "\<lbrakk>prj\<^sub>1 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\<rbrakk> = \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>1\<rbrakk>"
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.\<rho>\<sigma>.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
+ CMP_def SPN.FF_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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
+ using assms VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C 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 \<a> \<i> 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 \<lbrakk>\<lbrakk>rs.cmp\<rbrakk>\<rbrakk>"
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 \<Xi>: transformation_by_components VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<open>\<lambda>rs. CMP (fst rs) (snd rs)\<close>
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
+ using rs compositor_in_hom [of ?r ?s] VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
next
fix \<mu>\<nu>
assume \<mu>\<nu>: "VV.arr \<mu>\<nu>"
let ?\<mu> = "fst \<mu>\<nu>"
let ?\<nu> = "snd \<mu>\<nu>"
show "CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>)) \<bullet> HoSPN_SPN.map \<mu>\<nu> =
SPNoH.map \<mu>\<nu> \<bullet> CMP (fst (VV.dom \<mu>\<nu>)) (snd (VV.dom \<mu>\<nu>))"
proof -
let ?LHS = "CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>)) \<bullet> HoSPN_SPN.map \<mu>\<nu>"
let ?RHS = "SPNoH.map \<mu>\<nu> \<bullet> CMP (fst (VV.dom \<mu>\<nu>)) (snd (VV.dom \<mu>\<nu>))"
have LHS:
"Span.in_hom ?LHS (HoSPN_SPN.map (VV.dom \<mu>\<nu>)) (SPNoH.map (VV.cod \<mu>\<nu>))"
proof
show "Span.in_hom (HoSPN_SPN.map \<mu>\<nu>) (HoSPN_SPN.map (VV.dom \<mu>\<nu>))
(HoSPN_SPN.map (VV.cod \<mu>\<nu>))"
using \<mu>\<nu> by blast
show "Span.in_hom (CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>)))
(HoSPN_SPN.map (VV.cod \<mu>\<nu>)) (SPNoH.map (VV.cod \<mu>\<nu>))"
- using \<mu>\<nu> VV.cod_simp by (auto simp add: VV.arr_char)
+ using \<mu>\<nu> VV.cod_simp by (auto simp add: VV.arr_char\<^sub>S\<^sub>b\<^sub>C)
qed
have RHS:
"Span.in_hom ?RHS (HoSPN_SPN.map (VV.dom \<mu>\<nu>)) (SPNoH.map (VV.cod \<mu>\<nu>))"
- using \<mu>\<nu> VV.dom_simp VV.cod_simp by (auto simp add: VV.arr_char)
+ using \<mu>\<nu> VV.dom_simp VV.cod_simp by (auto simp add: VV.arr_char\<^sub>S\<^sub>b\<^sub>C)
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 \<mu>\<nu>)"
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 \<mu>\<nu>)"
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_\<mu>_\<nu>: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>dom ?\<mu>\<close> \<open>dom ?\<nu>\<close>
- using \<mu>\<nu> VV.ide_char VV.arr_char by unfold_locales auto
+ using \<mu>\<nu> VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by unfold_locales auto
interpret cod_\<mu>_\<nu>: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>cod ?\<mu>\<close> \<open>cod ?\<nu>\<close>
- using \<mu>\<nu> VV.ide_char VV.arr_char by unfold_locales auto
+ using \<mu>\<nu> VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by unfold_locales auto
interpret \<mu>_\<nu>: horizontal_composite_of_arrows_of_tabulations_in_maps
V H \<a> \<i> src trg
\<open>dom ?\<mu>\<close> \<open>tab_of_ide (dom ?\<mu>)\<close> \<open>tab\<^sub>0 (dom ?\<mu>)\<close> \<open>tab\<^sub>1 (dom ?\<mu>)\<close>
\<open>dom ?\<nu>\<close> \<open>tab_of_ide (dom ?\<nu>)\<close> \<open>tab\<^sub>0 (dom ?\<nu>)\<close> \<open>tab\<^sub>1 (dom ?\<nu>)\<close>
\<open>cod ?\<mu>\<close> \<open>tab_of_ide (cod ?\<mu>)\<close> \<open>tab\<^sub>0 (cod ?\<mu>)\<close> \<open>tab\<^sub>1 (cod ?\<mu>)\<close>
\<open>cod ?\<nu>\<close> \<open>tab_of_ide (cod ?\<nu>)\<close> \<open>tab\<^sub>0 (cod ?\<nu>)\<close> \<open>tab\<^sub>1 (cod ?\<nu>)\<close>
?\<mu> ?\<nu>
- using \<mu>\<nu> VV.arr_char by unfold_locales auto
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C by unfold_locales auto
let ?\<mu>\<nu> = "?\<mu> \<star> ?\<nu>"
interpret dom_\<mu>\<nu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom ?\<mu>\<nu>\<close>
using \<mu>\<nu> by (unfold_locales, simp)
interpret cod_\<mu>\<nu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>cod ?\<mu>\<nu>\<close>
using \<mu>\<nu> by (unfold_locales, simp)
interpret \<mu>\<nu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close> \<open>tab_of_ide (dom ?\<mu>\<nu>)\<close> \<open>tab\<^sub>0 (dom ?\<mu>\<nu>)\<close> \<open>tab\<^sub>1 (dom ?\<mu>\<nu>)\<close>
\<open>cod ?\<mu>\<nu>\<close> \<open>tab_of_ide (cod ?\<mu>\<nu>)\<close> \<open>tab\<^sub>0 (cod ?\<mu>\<nu>)\<close> \<open>tab\<^sub>1 (cod ?\<mu>\<nu>)\<close>
?\<mu>\<nu>
using \<mu>\<nu> by unfold_locales auto
have Chn_LHS_eq:
"Chn ?LHS = \<lbrakk>\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>\<rbrakk> \<odot> Span.chine_hcomp (SPN (fst \<mu>\<nu>)) (SPN (snd \<mu>\<nu>))"
proof -
have "Chn ?LHS = Chn (CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>))) \<odot>
Chn (HoSPN_SPN.map \<mu>\<nu>)"
using \<mu>\<nu> LHS Span.Chn_vcomp by blast
also have "... = \<lbrakk>\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>\<rbrakk> \<odot> Chn (HoSPN_SPN.map \<mu>\<nu>)"
- using \<mu>\<nu> VV.arr_char VV.cod_char CMP_def by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C CMP_def by simp
also have "... = \<lbrakk>\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>\<rbrakk> \<odot>
Span.chine_hcomp (SPN (fst \<mu>\<nu>)) (SPN (snd \<mu>\<nu>))"
- using \<mu>\<nu> VV.arr_char SPN.FF_def Span.hcomp_def by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<odot>
Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
proof -
have "Chn ?RHS = Chn (SPN (?\<mu> \<star> ?\<nu>)) \<odot>
Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
- using \<mu>\<nu> RHS Span.vcomp_def VV.arr_char CMP_def Span.arr_char Span.not_arr_Null
+ using \<mu>\<nu> RHS Span.vcomp_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C CMP_def Span.arr_char Span.not_arr_Null
VV.dom_simp
by auto
moreover have "Chn (SPN (?\<mu> \<star> ?\<nu>)) =
Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
proof -
have "Chn (SPN (?\<mu> \<star> ?\<nu>)) =
Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>spn ?\<mu>\<nu>\<rbrakk>"
using \<mu>\<nu> SPN_def by simp
also have "... = Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
using spn_def by simp
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
let ?Chn_LHS =
"Maps.MkArr (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<odot>
Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)"
let ?Chn_RHS =
"Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>))) (src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<odot>
Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
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 \<open>
Here is where we use \<open>dom_\<mu>_\<nu>.chine_hcomp_SPN_SPN\<close>,
which depends on our having chosen the ``right'' pullbacks for \<open>Maps(B)\<close>.
The map \<open>Chn_LHS\<close> has as its domain the apex of the
horizontal composite of the components of @{term "VV.dom \<mu>\<nu>"},
whereas \<open>Chn_RHS\<close> has as its
domain the apex of the chosen tabulation of \<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close>.
We need these to be equal in order for \<open>Chn_LHS\<close> and \<open>Chn_RHS\<close> to be equal.
\<close>
show "Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS"
proof -
have "Maps.Dom ?Chn_LHS = Maps.Dom (Maps.dom ?Chn_LHS)"
using \<mu>\<nu> 1 Maps.Dom_dom by presburger
also have
"... = Maps.Dom (Span.chine_hcomp (SPN (dom ?\<mu>)) (SPN (dom ?\<nu>)))"
proof -
have "... = Maps.Dom (Maps.dom (Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)))"
using 1 Maps.seq_char Maps.Dom_comp by auto
also have "... = Maps.Dom (Maps.pbdom (Leg0 (Dom (SPN ?\<mu>)))
(Leg1 (Dom (SPN ?\<nu>))))"
- using \<mu>\<nu> VV.arr_char Span.chine_hcomp_in_hom [of "SPN ?\<nu>" "SPN ?\<mu>"]
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C Span.chine_hcomp_in_hom [of "SPN ?\<nu>" "SPN ?\<mu>"]
by auto
also have "... = Maps.Dom (Maps.dom (Maps.pbdom (Leg0 (Dom (SPN ?\<mu>)))
(Leg1 (Dom (SPN ?\<nu>)))))"
proof -
have "Maps.cospan (Leg0 (Dom (SPN (fst \<mu>\<nu>)))) (Leg1 (Dom (SPN (snd \<mu>\<nu>))))"
- using \<mu>\<nu> VV.arr_char SPN_in_hom Span.arr_char Span.dom_char SPN_def
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C SPN_in_hom Span.arr_char Span.dom_char SPN_def
Maps.CLS_in_hom Maps.arr_char Maps.cod_char dom_\<mu>_\<nu>.composable
dom_\<mu>_\<nu>.RS_simps(16) dom_\<mu>_\<nu>.S\<^sub>1_def dom_\<mu>_\<nu>.RS_simps(1)
dom_\<mu>_\<nu>.R\<^sub>0_def Maps.pbdom_in_hom
by simp
thus ?thesis
- using \<mu>\<nu> VV.arr_char Maps.pbdom_in_hom by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C Maps.pbdom_in_hom by simp
qed
also have "... = Maps.Dom
(Maps.dom (Maps.pbdom (Leg0 (Dom (SPN (dom ?\<mu>))))
(Leg1 (Dom (SPN (dom ?\<nu>))))))"
- using \<mu>\<nu> SPN_def VV.arr_char by simp
+ using \<mu>\<nu> SPN_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = Maps.Dom
(Maps.dom (Span.chine_hcomp (SPN (dom ?\<mu>)) (SPN (dom ?\<nu>))))"
- using \<mu>\<nu> VV.arr_char ide_dom
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C ide_dom
by (simp add: Span.chine_hcomp_ide_ide)
also have "... = Maps.Dom (Span.chine_hcomp (SPN (dom ?\<mu>)) (SPN (dom ?\<nu>)))"
using Maps.Dom_dom Maps.in_homE SPN.preserves_reflects_arr SPN.preserves_src
SPN.preserves_trg Span.chine_hcomp_in_hom dom_\<mu>_\<nu>.composable
dom_\<mu>_\<nu>.r.base_simps(2) dom_\<mu>_\<nu>.s.base_simps(2)
by (metis (no_types, lifting))
finally show ?thesis by simp
qed
also have "... = src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0"
using "dom_\<mu>_\<nu>.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 ?\<mu> \<star> cod ?\<nu>))"
- using \<mu>\<nu> 1 VV.arr_char Maps.seq_char by auto
+ using \<mu>\<nu> 1 VV.arr_char\<^sub>S\<^sub>b\<^sub>C Maps.seq_char by auto
also have "... = src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>))"
- using \<mu>\<nu> VV.arr_char cod_\<mu>\<nu>.tab_simps(2) by auto
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C cod_\<mu>\<nu>.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 (\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp)"
proof -
have "Maps.Map ?Chn_RHS = Maps.Comp \<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
- using \<mu>\<nu> 2 VV.arr_char Maps.Map_comp
+ using \<mu>\<nu> 2 VV.arr_char\<^sub>S\<^sub>b\<^sub>C Maps.Map_comp
Maps.comp_char
[of "Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
"Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0)
(src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"]
by simp
also have "... = \<lbrakk>\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp\<rbrakk>"
proof -
have "\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk> \<in>
Maps.Hom (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))"
proof -
have "\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk> \<in>
Maps.Hom (src dom_\<mu>_\<nu>.\<rho>\<sigma>.tab) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))"
- using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.cmp_props(1-3)
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_\<mu>_\<nu>.cmp_props(1-3)
by (metis (mono_tags, lifting) equivalence_is_left_adjoint mem_Collect_eq)
thus ?thesis
- using \<mu>\<nu> VV.arr_char dom_\<mu>\<nu>.tab_simps(2) by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_\<mu>\<nu>.tab_simps(2) by simp
qed
moreover have "\<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<in>
Maps.Hom (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))"
- using \<mu>\<nu> VV.arr_char \<mu>\<nu>.chine_in_hom \<mu>\<nu>.is_map by auto
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<mu>\<nu>.chine_in_hom \<mu>\<nu>.is_map by auto
moreover have
"\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<in> Maps.Comp \<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
proof
show "is_iso_class \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
using is_iso_classI by simp
show "is_iso_class \<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
using is_iso_classI by simp
show "dom_\<mu>_\<nu>.cmp \<in> \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
using ide_in_iso_class [of dom_\<mu>_\<nu>.cmp] by simp
show "\<mu>\<nu>.chine \<in> \<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
using ide_in_iso_class by simp
show "\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<cong> \<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp"
- using \<mu>\<nu> VV.arr_char \<mu>\<nu>.chine_simps dom_\<mu>_\<nu>.cmp_simps dom_\<mu>\<nu>.tab_simps(2)
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<mu>\<nu>.chine_simps dom_\<mu>_\<nu>.cmp_simps dom_\<mu>\<nu>.tab_simps(2)
isomorphic_reflexive
by auto
qed
ultimately show ?thesis
using \<mu>\<nu> dom_\<mu>_\<nu>.cmp_props \<mu>\<nu>.chine_in_hom \<mu>\<nu>.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 = \<lbrakk>cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine\<rbrakk>"
proof -
have "Maps.Map ?Chn_LHS =
Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>
(Maps.Map
(Maps.tuple (Maps.CLS (spn ?\<mu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1))
(Maps.CLS (tab\<^sub>0 (cod ?\<mu>)))
(Maps.CLS (tab\<^sub>1 (cod ?\<nu>)))
(Maps.CLS (spn ?\<nu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0))))"
proof -
have "Maps.Map ?Chn_LHS =
Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>
(Maps.Map (Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)))"
- using \<mu>\<nu> 1 VV.arr_char Maps.Map_comp cod_\<mu>\<nu>.tab_simps(2)
+ using \<mu>\<nu> 1 VV.arr_char\<^sub>S\<^sub>b\<^sub>C Maps.Map_comp cod_\<mu>\<nu>.tab_simps(2)
Maps.comp_char
[of "Maps.MkArr (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0)
(src (tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>"
"Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)"]
by simp
moreover have "Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>) =
Maps.tuple
(Maps.CLS (spn ?\<mu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1))
(Maps.CLS (tab\<^sub>0 (cod ?\<mu>)))
(Maps.CLS (tab\<^sub>1 (cod ?\<nu>)))
(Maps.CLS (spn ?\<nu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0))"
proof -
have "Maps.PRJ\<^sub>0
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<and>
Maps.PRJ\<^sub>1
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
interpret X: identity_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>(tab\<^sub>0 (dom ?\<mu>))\<^sup>* \<star> tab\<^sub>1 (dom ?\<nu>)\<close>
- using \<mu>\<nu> VV.arr_char
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (unfold_locales, simp)
have "Maps.PRJ\<^sub>0
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg (snd \<mu>\<nu>))
\<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>)
\<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>))\<rbrakk>\<rbrakk>"
unfolding Maps.PRJ\<^sub>0_def
- using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.RS_simps(1) dom_\<mu>_\<nu>.RS_simps(16)
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_\<mu>_\<nu>.RS_simps(1) dom_\<mu>_\<nu>.RS_simps(16)
dom_\<mu>_\<nu>.RS_simps(18) dom_\<mu>_\<nu>.RS_simps(3) dom_\<mu>_\<nu>.R\<^sub>0_def
dom_\<mu>_\<nu>.S\<^sub>1_def
by auto
moreover
have "Maps.PRJ\<^sub>1
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>tab\<^sub>1 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg (snd \<mu>\<nu>))
\<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>)
\<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>))\<rbrakk>\<rbrakk>"
unfolding Maps.PRJ\<^sub>1_def
- using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.RS_simps(1) dom_\<mu>_\<nu>.RS_simps(16)
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_\<mu>_\<nu>.RS_simps(1) dom_\<mu>_\<nu>.RS_simps(16)
dom_\<mu>_\<nu>.RS_simps(18) dom_\<mu>_\<nu>.RS_simps(3) dom_\<mu>_\<nu>.R\<^sub>0_def
dom_\<mu>_\<nu>.S\<^sub>1_def
by auto
moreover
have "(Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg (snd \<mu>\<nu>))
\<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>)
\<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) \<cong>
(tab\<^sub>0 (dom ?\<mu>))\<^sup>* \<star> tab\<^sub>1 (dom ?\<nu>)"
- using VV.arr_char \<mu>\<nu> dom_\<mu>_\<nu>.S\<^sub>1_def dom_\<mu>_\<nu>.s.leg1_simps(3)
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<mu>\<nu> dom_\<mu>_\<nu>.S\<^sub>1_def dom_\<mu>_\<nu>.s.leg1_simps(3)
dom_\<mu>_\<nu>.s.leg1_simps(4) trg_dom dom_\<mu>_\<nu>.R\<^sub>0_def
dom_\<mu>_\<nu>.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 \<mu>\<nu> VV.arr_char SPN_def isomorphic_reflexive
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C SPN_def isomorphic_reflexive
Maps.comp_CLS [of "spn ?\<mu>" dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1 "spn ?\<mu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1"]
Maps.comp_CLS [of "spn ?\<nu>" dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0 "spn ?\<nu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0"]
by simp
qed
moreover have "Maps.Dom (Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)) =
src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0"
by (metis (no_types, lifting) "1" "2" Maps.Dom.simps(1) Maps.comp_char
\<open>Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS\<close>)
ultimately show ?thesis by simp
qed
also have "... = Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
proof -
let ?tuple = "Maps.tuple \<lbrakk>\<lbrakk>spn (fst \<mu>\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 (cod ?\<mu>)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 (cod ?\<nu>)\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>spn (snd \<mu>\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
have "iso_class \<mu>_\<nu>.chine = Maps.Map ?tuple"
using \<mu>_\<nu>.CLS_chine spn_def Maps.Map.simps(1)
by (metis (no_types, lifting))
thus ?thesis by simp
qed
also have "... = \<lbrakk>cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine\<rbrakk>"
proof -
have "\<lbrakk>\<mu>_\<nu>.chine\<rbrakk> \<in> Maps.Hom (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0)"
proof -
have "\<guillemotleft>\<mu>_\<nu>.chine : src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0 \<rightarrow> src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<guillemotright>"
- using \<mu>\<nu> VV.arr_char by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
thus ?thesis
using \<mu>_\<nu>.is_map ide_in_iso_class left_adjoint_is_ide by blast
qed
moreover have "\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<in>
Maps.Hom (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))"
proof -
have "\<guillemotleft>cod_\<mu>_\<nu>.cmp : src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0 \<rightarrow> src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>))\<guillemotright>"
- using \<mu>\<nu> VV.arr_char cod_\<mu>_\<nu>.cmp_in_hom cod_\<mu>\<nu>.tab_simps(2)
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C cod_\<mu>_\<nu>.cmp_in_hom cod_\<mu>\<nu>.tab_simps(2)
by simp
thus ?thesis
using cod_\<mu>_\<nu>.cmp_props equivalence_is_left_adjoint left_adjoint_is_ide
ide_in_iso_class [of cod_\<mu>_\<nu>.cmp]
by (metis (mono_tags, lifting) mem_Collect_eq)
qed
moreover have
"cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine \<in> Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
proof
show "is_iso_class \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
using \<mu>_\<nu>.w_simps(1) is_iso_classI by blast
show "is_iso_class \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>"
using cod_\<mu>_\<nu>.cmp_simps(2) is_iso_classI by blast
show "\<mu>_\<nu>.chine \<in> \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
using ide_in_iso_class by simp
show "cod_\<mu>_\<nu>.cmp \<in> \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>"
using ide_in_iso_class by simp
show "cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine \<cong> cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine"
by (simp add: isomorphic_reflexive)
qed
ultimately show ?thesis
using \<mu>\<nu> cod_\<mu>_\<nu>.cmp_props \<mu>_\<nu>.chine_in_hom \<mu>_\<nu>.chine_is_induced_map
Maps.Comp_eq_iso_class_memb
by simp
qed
finally show ?thesis by simp
qed
have EQ: "\<lbrakk>\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp\<rbrakk> = \<lbrakk>cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine\<rbrakk>"
proof (intro iso_class_eqI)
show "\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<cong> cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine"
proof -
interpret dom_cmp: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close>
dom_\<mu>_\<nu>.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 (dom ?\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (dom ?\<mu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>dom ?\<mu>\<nu>\<close>
\<open>tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)\<close>
\<open>tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)\<close>
\<open>tab\<^sub>1 (dom ?\<mu> \<star> dom ?\<nu>)\<close>
\<open>dom ?\<mu>\<nu>\<close>
- using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.cmp_interpretation by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_\<mu>_\<nu>.cmp_interpretation by simp
interpret cod_cmp: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>cod ?\<mu>\<nu>\<close>
cod_\<mu>_\<nu>.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 (cod ?\<nu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (cod ?\<mu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>cod ?\<mu>\<nu>\<close>
\<open>tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>)\<close>
\<open>tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)\<close>
\<open>tab\<^sub>1 (cod ?\<mu> \<star> cod ?\<nu>)\<close>
\<open>cod ?\<mu>\<nu>\<close>
- using \<mu>\<nu> VV.arr_char cod_\<mu>_\<nu>.cmp_interpretation by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C cod_\<mu>_\<nu>.cmp_interpretation by simp
interpret L: vertical_composite_of_arrows_of_tabulations_in_maps
V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close>
\<open>dom_\<mu>_\<nu>.\<rho>\<sigma>.tab\<close>
\<open>tab\<^sub>0 (dom ?\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (dom ?\<mu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>dom ?\<mu>\<nu>\<close>
\<open>tab_of_ide (dom ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>0 (dom ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>1 (dom ?\<mu>\<nu>)\<close>
\<open>cod ?\<mu>\<nu>\<close>
cod_\<mu>\<nu>.tab
\<open>tab\<^sub>0 (cod ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>1 (cod ?\<mu>\<nu>)\<close>
\<open>dom ?\<mu>\<nu>\<close>
\<open>?\<mu> \<star> ?\<nu>\<close>
- using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.cmp_in_hom
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_\<mu>_\<nu>.cmp_in_hom
by unfold_locales auto
interpret R: vertical_composite_of_arrows_of_tabulations_in_maps
V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close>
\<open>dom_\<mu>_\<nu>.\<rho>\<sigma>.tab\<close>
\<open>tab\<^sub>0 (dom ?\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (dom ?\<mu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>cod ?\<mu>\<nu>\<close>
\<open>cod_\<mu>_\<nu>.\<rho>\<sigma>.tab\<close>
\<open>tab\<^sub>0 (cod ?\<nu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (cod ?\<mu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>cod ?\<mu>\<nu>\<close>
cod_\<mu>\<nu>.tab
\<open>tab\<^sub>0 (cod ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>1 (cod ?\<mu>\<nu>)\<close>
\<open>?\<mu> \<star> ?\<nu>\<close>
\<open>cod ?\<mu>\<nu>\<close>
- using \<mu>\<nu> VV.arr_char cod_\<mu>_\<nu>.cmp_in_hom
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C cod_\<mu>_\<nu>.cmp_in_hom
by unfold_locales auto
have "\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<cong> L.chine"
- using \<mu>\<nu> VV.arr_char L.chine_char dom_\<mu>_\<nu>.cmp_def isomorphic_symmetric
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C L.chine_char dom_\<mu>_\<nu>.cmp_def isomorphic_symmetric
by simp
also have "... = R.chine"
using L.is_ide \<mu>\<nu> comp_arr_dom comp_cod_arr isomorphic_reflexive by force
also have "... \<cong> cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine"
- using \<mu>\<nu> VV.arr_char R.chine_char cod_\<mu>_\<nu>.cmp_def by simp
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C R.chine_char cod_\<mu>_\<nu>.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 \<Xi>: natural_isomorphism VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<Xi>.map
- using VV.ide_char VV.arr_char \<Xi>.map_simp_ide compositor_is_iso
+ using VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Xi>.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
(\<lambda>rs. CMP (fst rs) (snd rs))"
..
lemma compositor_is_natural_isomorphism:
shows "natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \<Xi>.map"
..
end
subsubsection "Associativity Coherence"
locale three_composable_identities_in_bicategory_of_spans =
bicategory_of_spans V H \<a> \<i> src trg +
f: identity_in_bicategory_of_spans V H \<a> \<i> src trg f +
g: identity_in_bicategory_of_spans V H \<a> \<i> src trg g +
h: identity_in_bicategory_of_spans V H \<a> \<i> src trg h
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> '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 \<a> \<i> src trg
f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f
using f.is_arrow_of_tabulations_in_maps by simp
interpretation h: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
h h.tab \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close> h h.tab \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close> h
using h.is_arrow_of_tabulations_in_maps by simp
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 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 \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<circ> snd \<mu>\<nu>\<close>
..
interpretation SPNoH: composite_functor VV.comp V Span.vcomp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> SPN
..
text \<open>
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 \<open>H\<close> refers to horizontal
composition of 1-cells and \<open>T\<close> refers to composite of tabulations.
So, for example, \<open>THfgh\<close> refers to the composite of the tabulation associated with the
horizontal composition \<open>f \<star> g\<close> with the tabulation associated with \<open>h\<close>.
\<close>
interpretation HHfgh: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>(f \<star> g) \<star> h\<close>
using fg gh by unfold_locales auto
interpretation HfHgh: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>f \<star> g \<star> h\<close>
using fg gh by unfold_locales auto
interpretation Tfg: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg f g
using fg gh by unfold_locales auto
interpretation Tgh: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg g h
using fg gh by unfold_locales auto
interpretation THfgh: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>f \<star> g\<close> h
using fg gh by unfold_locales auto
interpretation THfgh: tabulation V H \<a> \<i> src trg \<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
using THfgh.\<rho>\<sigma>.composite_is_tabulation by simp
interpretation TfHgh: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
f \<open>g \<star> h\<close>
using fg gh by unfold_locales auto
interpretation TfHgh: tabulation V H \<a> \<i> src trg \<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
using TfHgh.\<rho>\<sigma>.composite_is_tabulation by simp
interpretation Tfg_Hfg: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>f \<star> g\<close> Tfg.\<rho>\<sigma>.tab \<open>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g\<close> \<open>tab_of_ide (f \<star> g)\<close> \<open>tab\<^sub>0 (f \<star> g)\<close> \<open>tab\<^sub>1 (f \<star> g)\<close>
\<open>f \<star> g\<close>
by unfold_locales auto
interpretation Tgh_Hgh: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>g \<star> h\<close> \<open>tab_of_ide (g \<star> h)\<close> \<open>tab\<^sub>0 (g \<star> h)\<close> \<open>tab\<^sub>1 (g \<star> h)\<close>
\<open>g \<star> h\<close>
by unfold_locales auto
interpretation THfgh_HHfgh:
arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>(f \<star> g) \<star> h\<close> \<open>tab_of_ide ((f \<star> g) \<star> h)\<close> \<open>tab\<^sub>0 ((f \<star> g) \<star> h)\<close> \<open>tab\<^sub>1 ((f \<star> g) \<star> h)\<close>
\<open>(f \<star> g) \<star> h\<close>
using fg gh by unfold_locales auto
interpretation TfHgh_HfHgh:
arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> \<open>tab_of_ide (f \<star> g \<star> h)\<close> \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>f \<star> g \<star> h\<close>
using fg gh by unfold_locales auto
interpretation TTfgh: composite_tabulation_in_maps V H \<a> \<i> src trg
\<open>f \<star> g\<close> Tfg.\<rho>\<sigma>.tab \<open>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1\<close>
h \<open>tab_of_ide h\<close> \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close>
using fg gh by unfold_locales auto
interpretation TTfgh_THfgh:
horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>f \<star> g\<close> Tfg.\<rho>\<sigma>.tab \<open>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1\<close>
h \<open>tab_of_ide h\<close> \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close>
\<open>f \<star> g\<close> \<open>tab_of_ide (f \<star> g)\<close> \<open>tab\<^sub>0 (f \<star> g)\<close> \<open>tab\<^sub>1 (f \<star> g)\<close>
h \<open>tab_of_ide h\<close> \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close>
\<open>f \<star> g\<close> h
..
interpretation TfTgh: composite_tabulation_in_maps V H \<a> \<i> src trg
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
using fg gh by unfold_locales auto
interpretation TfTgh_TfHgh:
horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> \<open>tab_of_ide (g \<star> h)\<close> \<open>tab\<^sub>0 (g \<star> h)\<close> \<open>tab\<^sub>1 (g \<star> h)\<close>
f \<open>g \<star> h\<close>
..
interpretation TfTgh_TfTgh:
horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
f \<open>g \<star> h\<close>
..
text \<open>
The following interpretation defines the associativity between the peaks
of the two composite tabulations \<open>TTfgh\<close> (associated to the left) and \<open>TfTgh\<close>
(associated to the right).
\<close>
(* TODO: Try to get rid of the .\<rho>\<sigma> in, e.g., Tfg.\<rho>\<sigma>.p\<^sub>1. *)
interpretation TTfgh_TfTgh:
arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfTgh.tab \<open>(tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0) \<star> TfTgh.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfTgh.p\<^sub>1\<close>
\<open>\<a>[f, g, h]\<close>
using fg gh by unfold_locales auto
text \<open>
This interpretation defines the map, from the apex of the tabulation associated
with the horizontal composite \<open>(f \<star> g) \<star> h\<close> to the apex of the tabulation associated
with the horizontal composite \<open>f \<star> g \<star> h\<close>, induced by the associativity isomorphism
\<open>\<a>[f, g, h]\<close> from \<open>(f \<star> g) \<star> h\<close> to \<open>f \<star> g \<star> h\<close>.
\<close>
interpretation HHfgh_HfHgh: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom (\<alpha> (f, g, h))\<close> \<open>tab_of_ide (dom (\<alpha> (f, g, h)))\<close>
\<open>tab\<^sub>0 (dom (\<alpha> (f, g, h)))\<close> \<open>tab\<^sub>1 (dom (\<alpha> (f, g, h)))\<close>
\<open>cod (\<alpha> (f, g, h))\<close> \<open>tab_of_ide (cod (\<alpha> (f, g, h)))\<close>
\<open>tab\<^sub>0 (cod (\<alpha> (f, g, h)))\<close> \<open>tab\<^sub>1 (cod (\<alpha> (f, g, h)))\<close>
\<open>\<alpha> (f, g, h)\<close>
proof -
have "arrow_of_tabulations_in_maps V H \<a> \<i> src trg
((f \<star> g) \<star> h) (tab_of_ide ((f \<star> g) \<star> h)) (tab\<^sub>0 ((f \<star> g) \<star> h)) (tab\<^sub>1 ((f \<star> g) \<star> h))
(f \<star> g \<star> h) (tab_of_ide (f \<star> g \<star> h)) (tab\<^sub>0 (f \<star> g \<star> h)) (tab\<^sub>1 (f \<star> g \<star> h))
\<a>[f, g, h]"
using fg gh by unfold_locales auto
thus "arrow_of_tabulations_in_maps V H \<a> \<i> src trg
(dom (\<alpha> (f, g, h))) (tab_of_ide (dom (\<alpha> (f, g, h))))
(tab\<^sub>0 (dom (\<alpha> (f, g, h)))) (tab\<^sub>1 (dom (\<alpha> (f, g, h))))
(cod (\<alpha> (f, g, h))) (tab_of_ide (cod (\<alpha> (f, g, h))))
(tab\<^sub>0 (cod (\<alpha> (f, g, h)))) (tab\<^sub>1 (cod (\<alpha> (f, g, h))))
(\<alpha> (f, g, h))"
using fg gh \<alpha>_def by auto
qed
interpretation SPN_f: arrow_of_spans Maps.comp \<open>SPN f\<close>
using SPN_in_hom Span.arr_char [of "SPN f"] by simp
interpretation SPN_g: arrow_of_spans Maps.comp \<open>SPN g\<close>
using SPN_in_hom Span.arr_char [of "SPN g"] by simp
interpretation SPN_h: arrow_of_spans Maps.comp \<open>SPN h\<close>
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 \<open>SPN f\<close> \<open>SPN g\<close> \<open>SPN h\<close>
using fg gh Span.arr_char SPN_in_hom SPN.preserves_ide Span.ide_char
apply unfold_locales by auto
text \<open>
The following relates the projections associated with the composite span \<open>SPN_fgh\<close>
with tabulations in the underlying bicategory.
\<close>
lemma prj_char:
shows "SPN_fgh.Prj\<^sub>1\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>0\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>0 = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>1 = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>1\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>1 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>0\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>0 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
show "SPN_fgh.Prj\<^sub>1\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "ide (Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
by (metis TTfgh.composable TTfgh.leg1_simps(2) Tfg.\<rho>\<sigma>.T0.antipar(2)
Tfg.\<rho>\<sigma>.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.\<rho>\<sigma>.prj_char TTfgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\<rho>\<sigma>.p\<^sub>0 "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
Maps.comp_CLS [of Tfg.\<rho>\<sigma>.p\<^sub>1 TTfgh.p\<^sub>1 "Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1"]
by (simp add: TTfgh.composable Tfg.\<rho>\<sigma>.T0.antipar(2))
qed
show "SPN_fgh.Prj\<^sub>0\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "ide (Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> 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.\<rho>\<sigma>.T0.antipar(2) Tfg.\<rho>\<sigma>.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.\<rho>\<sigma>.prj_char TTfgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\<rho>\<sigma>.p\<^sub>0 "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
Maps.comp_CLS [of Tfg.\<rho>\<sigma>.p\<^sub>0 TTfgh.p\<^sub>1 "Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1"]
by (simp add: Tfg.\<rho>\<sigma>.T0.antipar(2) THfgh.composable)
qed
show "SPN_fgh.Prj\<^sub>0 = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using isomorphic_reflexive TTfgh.prj_char Tfg.\<rho>\<sigma>.prj_char
Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\<rho>\<sigma>.p\<^sub>0 "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
by (simp add: Tfg.composable)
show "SPN_fgh.Prj\<^sub>1 = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using Tgh.\<rho>\<sigma>.prj_char isomorphic_reflexive Tgh.composable
Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\<rho>\<sigma>.p\<^sub>1 "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
TfTgh.prj_char
by simp
show "SPN_fgh.Prj\<^sub>1\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>1 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.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 \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
Maps.comp_CLS [of Tgh.\<rho>\<sigma>.p\<^sub>1 TfTgh.p\<^sub>0 "Tgh.\<rho>\<sigma>.p\<^sub>1 \<star> TfTgh.p\<^sub>0"]
by simp
show "SPN_fgh.Prj\<^sub>0\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>0 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.prj_char TfTgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" "Tgh.\<rho>\<sigma>.p\<^sub>1" "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
Maps.comp_CLS [of Tgh.\<rho>\<sigma>.p\<^sub>0 TfTgh.p\<^sub>0 "Tgh.\<rho>\<sigma>.p\<^sub>0 \<star> TfTgh.p\<^sub>0"]
by simp
qed
interpretation \<Phi>: transformation_by_components VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<open>\<lambda>rs. CMP (fst rs) (snd rs)\<close>
using compositor_is_natural_transformation by simp
interpretation \<Phi>: natural_isomorphism VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<Phi>.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
\<open>\<lambda>\<tau>\<mu>\<nu>. arr (fst \<tau>\<mu>\<nu>) \<and> arr (fst (snd \<tau>\<mu>\<nu>)) \<and> arr (snd (snd \<tau>\<mu>\<nu>)) \<and>
src (fst (snd \<tau>\<mu>\<nu>)) = trg (snd (snd \<tau>\<mu>\<nu>)) \<and>
src (fst \<tau>\<mu>\<nu>) = trg (fst (snd \<tau>\<mu>\<nu>))\<close>
- using fg gh VVV.arr_char VV.arr_char VVV.subcategory_axioms by simp
+ using fg gh VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.subcategory_axioms by simp
text \<open>
We define abbreviations for the left and right-hand sides of the equation for
associativity coherence.
\<close>
(*
* TODO: \<Phi> doesn't really belong in this locale. Replace it with CMP and rearrange
* material so that this locale comes first and the definition of \<Phi> comes later
* in bicategory_of_spans.
*)
abbreviation LHS
where "LHS \<equiv> SPN \<a>[f, g, h] \<bullet> \<Phi>.map (f \<star> g, h) \<bullet> (\<Phi>.map (f, g) \<circ> SPN h)"
abbreviation RHS
where "RHS \<equiv> \<Phi>.map (f, g \<star> h) \<bullet> (SPN f \<circ> \<Phi>.map (g, h)) \<bullet>
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'
+ using fg gh VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C Span.hseqI'
HoHV_def compositor_in_hom \<alpha>_def
apply (intro Span.seqI)
apply simp_all
using SPN.FF_def
apply simp
proof -
have "SPN ((f \<star> g) \<star> h) = Span.cod (CMP (f \<star> g) h)"
using fg gh compositor_in_hom by simp
also have "... = Span.cod (CMP (f \<star> g) h \<bullet> (CMP f g \<circ> SPN h))"
proof -
have "Span.seq (CMP (f \<star> g) h) (CMP f g \<circ> 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.\<nu>\<pi>.composable fg by auto
show 3: "Span.arr (CMP (f \<star> g) h)"
using TTfgh.composable Tfg.\<rho>\<sigma>.ide_base compositor_simps(1) h.is_ide by auto
show "Span.dom (CMP (f \<star> g) h) = Span.cod (CMP f g \<circ> SPN h)"
using 1 2 3 fg gh compositor_in_hom SPN_fgh.\<nu>\<pi>.composable SPN_in_hom SPN.FF_def
by auto
qed
thus ?thesis by simp
qed
finally show "SPN ((f \<star> g) \<star> h) = Span.cod (CMP (f \<star> g) h \<bullet> (CMP f g \<circ> SPN h))"
by blast
qed
lemma arr_RHS:
shows "Span.arr RHS"
- using fg gh VV.ide_char VV.arr_char \<Phi>.map_simp_ide SPN.FF_def Span.hseqI'
+ using fg gh VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>.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 (\<Phi>.map (f, g) \<circ> SPN h)"
using arr_LHS by auto
also have "... = Span.dom (\<Phi>.map (f, g)) \<circ> Span.dom (SPN h)"
using arr_LHS Span.dom_hcomp [of "SPN h" "\<Phi>.map (f, g)"] by blast
also have "... = (SPN f \<circ> SPN g) \<circ> SPN h"
- using fg \<Phi>.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def by simp
+ using fg \<Phi>.map_simp_ide VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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
+ using fg gh VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = Span.dom RHS"
using \<open>Span.arr RHS\<close> by auto
finally show ?thesis by blast
qed
show "Span.cod LHS = Span.cod RHS"
proof -
have "Span.cod LHS = Span.cod (SPN \<a>[f, g, h])"
using arr_LHS by simp
also have "... = SPN (f \<star> g \<star> h)"
unfolding \<alpha>_def
- using fg gh VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char HoVH_def
+ using fg gh VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C HoVH_def
by simp
also have "... = Span.cod RHS"
- using arr_RHS fg gh \<Phi>.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def
+ using arr_RHS fg gh \<Phi>.map_simp_ide VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C SPN.FF_def
compositor_in_hom
by simp
finally show ?thesis by blast
qed
qed
lemma Chn_LHS_eq:
shows "Chn LHS =
\<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Chn LHS = \<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot>
Span.chine_hcomp (CMP f g) (SPN h)"
proof -
have "Chn LHS = Chn (SPN \<a>[f, g, h]) \<odot> Chn (CMP (f \<star> g) h) \<odot>
Chn (CMP f g \<circ> SPN h)"
- using fg gh arr_LHS \<Phi>.map_simp_ide VV.ide_char VV.arr_char Span.Chn_vcomp
+ using fg gh arr_LHS \<Phi>.map_simp_ide VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C Span.Chn_vcomp
by auto
moreover have "Chn (SPN \<a>[f, g, h]) = Maps.CLS HHfgh_HfHgh.chine"
- using fg gh SPN_def VVV.arr_char VV.arr_char spn_def \<alpha>_def by simp
+ using fg gh SPN_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C spn_def \<alpha>_def by simp
moreover have "Chn (CMP (f \<star> g) h) = Maps.CLS THfgh_HHfgh.chine"
using fg gh CMP_def THfgh.cmp_def by simp
moreover have "Chn (CMP f g \<circ> 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) = \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Span.chine_hcomp (CMP f g) (SPN h) =
Maps.tuple
(\<lbrakk>\<lbrakk>Tfg.cmp\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>spn h\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)"
proof -
have "\<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using fg isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>0 g" "Tfg.\<rho>\<sigma>.p\<^sub>0" "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
by simp
moreover have "span_in_category.apex Maps.comp
\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 h\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>\<rparr> =
\<lbrakk>\<lbrakk>spn h\<rbrakk>\<rbrakk>"
proof -
interpret h: span_in_category Maps.comp \<open>\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 h\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>\<rparr>\<close>
using h.determines_span by simp
interpret dom_h: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom h\<close> \<open>tab_of_ide (dom h)\<close> \<open>tab\<^sub>0 (dom h)\<close> \<open>tab\<^sub>1 (dom h)\<close>
\<open>cod h\<close> \<open>tab_of_ide (cod h)\<close> \<open>tab\<^sub>0 (cod h)\<close> \<open>tab\<^sub>1 (cod h)\<close>
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 = \<lbrakk>\<lbrakk>dom_h.chine\<rbrakk>\<rbrakk>"
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.\<Delta>_eq_\<rho> h.apex_is_induced_by_cell by force
hence "src (tab\<^sub>0 h) \<cong> h.chine"
using h.chine_is_induced_map h.induced_map_unique by simp
thus "\<lbrakk>src (tab\<^sub>0 h)\<rbrakk> = \<lbrakk>h.chine\<rbrakk>"
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.\<rho>\<sigma>.prj_char Span.hcomp_def by simp
qed
also have "... = \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "\<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk> =
Maps.tuple \<lbrakk>\<lbrakk>Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>h.chine \<star> TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using TTfgh_THfgh.CLS_chine by simp
also have "... =
Maps.tuple (\<lbrakk>\<lbrakk>Tfg_Hfg.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>h.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>)"
proof -
have "\<lbrakk>\<lbrakk>Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg_Hfg.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint TTfgh.p\<^sub>1"
using Tfg.\<rho>\<sigma>.T0.antipar(2) THfgh.composable by simp
moreover have "Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1 \<cong> Tfg_Hfg.chine \<star> 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 \<star> TTfgh.p\<^sub>1"]
by simp
qed
moreover have "\<lbrakk>\<lbrakk>h.chine \<star> TTfgh.p\<^sub>0\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>h.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint TTfgh.p\<^sub>0"
by (simp add: Tfg.\<rho>\<sigma>.T0.antipar(2) THfgh.composable)
moreover have "h.chine \<star> TTfgh.p\<^sub>0 \<cong> h.chine \<star> 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 \<star> TTfgh.p\<^sub>0"]
by simp
qed
ultimately show ?thesis by argo
qed
also have "... =
Maps.tuple (\<lbrakk>\<lbrakk>Tfg.cmp\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>spn h\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)"
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 \<equiv> Maps.tuple SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 SPN_fgh.Prj\<^sub>0"
abbreviation tuple_ABC
where "tuple_ABC \<equiv> Maps.tuple SPN_fgh.Prj\<^sub>1\<^sub>1
SPN_fgh.\<mu>.leg0
(SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
tuple_BC"
abbreviation tuple_BC'
where "tuple_BC' \<equiv> Maps.tuple \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
abbreviation tuple_ABC'
where "tuple_ABC' \<equiv> Maps.tuple \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>
tuple_BC'"
lemma csq:
shows "Maps.commutative_square SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1
SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
and "Maps.commutative_square SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
proof -
show 1: "Maps.commutative_square SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1
SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
proof
show "Maps.cospan SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1"
using SPN_fgh.\<nu>\<pi>.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.\<nu>.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.\<nu>.leg0 \<odot> SPN_fgh.Prj\<^sub>0\<^sub>1 = SPN_fgh.\<pi>.leg1 \<odot> SPN_fgh.Prj\<^sub>0"
by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc
Maps.pullback_commutes' SPN_fgh.\<mu>\<nu>.dom.leg_simps(1)
SPN_fgh.\<mu>\<nu>.leg0_composite SPN_fgh.cospan_\<nu>\<pi>)
qed
show "Maps.commutative_square
SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.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.\<mu>.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.\<mu>.leg0 \<odot> SPN_fgh.Prj\<^sub>1\<^sub>1 = (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_BC"
using 1 fg gh Maps.comp_assoc Maps.prj_tuple
by (metis (no_types, lifting) Maps.pullback_commutes' SPN_fgh.cospan_\<mu>\<nu>)
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 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using prj_char by simp
moreover have "SPN_fgh.\<mu>.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk>"
by simp
moreover have "SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1 = \<lbrakk>\<lbrakk>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
using Tgh.\<rho>\<sigma>.prj_char isomorphic_reflexive Tgh.composable
Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\<rho>\<sigma>.p\<^sub>1 "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
by (simp add: g.T0.antipar(2))
moreover show "tuple_BC = tuple_BC'"
using prj_char Tfg.\<rho>\<sigma>.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.\<rho>\<sigma>.p\<^sub>0))"
proof
show 1: "Maps.arr tuple_BC"
using csq(1) by simp
have 2: "Maps.commutative_square
\<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
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.\<rho>\<sigma>.p\<^sub>0)"
proof -
have "Maps.cod tuple_BC' = Maps.pbdom \<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>"
using 1 2 fg gh Maps.tuple_in_hom by blast
also have "... = Maps.MkIde (src Tgh.\<rho>\<sigma>.p\<^sub>0)"
using 1 2 fg gh Maps.pbdom_def
by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.\<nu>\<pi>.are_identities(2)
SPN_fgh.\<nu>\<pi>.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 \<circ> SPN g) \<circ> 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 \<circ> SPN g \<circ> 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 = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> tuple_ABC'"
proof -
have "Chn RHS =
Chn (\<Phi>.map (f, g \<star> h)) \<odot> Chn (SPN f \<circ> \<Phi>.map (g, h)) \<odot>
Chn (Span.assoc (SPN f) (SPN g) (SPN h))"
proof -
have "Chn RHS = Chn (\<Phi>.map (f, g \<star> h)) \<odot>
Chn ((SPN f \<circ> \<Phi>.map (g, h)) \<bullet> Span.assoc (SPN f) (SPN g) (SPN h))"
using arr_RHS Span.vcomp_eq Span.Chn_vcomp by blast
also have "... = Chn (\<Phi>.map (f, g \<star> h)) \<odot> Chn (SPN f \<circ> \<Phi>.map (g, h)) \<odot>
Chn (Span.assoc (SPN f) (SPN g) (SPN h))"
proof -
have "Span.seq (SPN f \<circ> \<Phi>.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 \<circ> \<Phi>.map (g, h)"
"Span.assoc (SPN f) (SPN g) (SPN h)"]
by simp
qed
finally show ?thesis by blast
qed
moreover have "Chn (\<Phi>.map (f, g \<star> h)) = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk>"
- using arr_RHS fg gh \<Phi>.map_simp_ide VV.ide_char VV.arr_char CMP_def TfHgh.cmp_def
+ using arr_RHS fg gh \<Phi>.map_simp_ide VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C CMP_def TfHgh.cmp_def
by simp
moreover have "Chn (SPN f \<circ> \<Phi>.map (g, h)) = Span.chine_hcomp (SPN f) (CMP g h)"
- using fg gh Span.hcomp_def \<Phi>.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def
+ using fg gh Span.hcomp_def \<Phi>.map_simp_ide VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C SPN.FF_def
by simp
moreover have "Chn (Span.assoc (SPN f) (SPN g) (SPN h)) = tuple_ABC"
- using fg gh Span.\<alpha>_ide VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char
+ using fg gh Span.\<alpha>_ide VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
SPN_fgh.chine_assoc_def Span.\<alpha>_def
by simp
moreover have "Span.chine_hcomp (SPN f) (CMP g h) = \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Span.chine_hcomp (SPN f) (CMP g h) =
Maps.tuple
(\<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 (g \<star> h)\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>Tgh_Hgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>)"
proof -
interpret f: span_in_category Maps.comp
\<open>\<lparr>Leg0 = Maps.MkArr (src (tab\<^sub>0 f)) (trg g) \<lbrakk>tab\<^sub>0 f\<rbrakk>,
Leg1 = Maps.MkArr (src (tab\<^sub>0 f)) (trg f) \<lbrakk>tab\<^sub>1 f\<rbrakk>\<rparr>\<close>
using f.determines_span
by (simp add: Tfg.composable)
interpret f: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> 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 \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
using Maps.CLS_in_hom f.is_map by blast
show "Maps.Dom f.apex = Maps.Dom \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto
show "Maps.Cod f.apex = Maps.Cod \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto
show "Maps.Map f.apex = Maps.Map \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
proof -
have "Maps.Map f.apex = \<lbrakk>src (tab\<^sub>0 f)\<rbrakk>"
using f.apex_def Maps.dom_char Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable
by auto
also have "... = \<lbrakk>f.chine\<rbrakk>"
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) \<cong> f.chine"
using f.induced_map_unique f.chine_is_induced_map by simp
qed
also have "... = Maps.Map \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
by simp
finally show ?thesis by simp
qed
qed
thus ?thesis
unfolding Span.chine_hcomp_def
using fg gh CMP_def Tgh.\<rho>\<sigma>.prj_char Span.hcomp_def isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\<rho>\<sigma>.p\<^sub>1 "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
Tgh.cmp_def TfTgh.prj_char
by simp
qed
also have "... = Maps.tuple \<lbrakk>\<lbrakk>f.chine \<star> TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 (g \<star> h)\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>Tgh_Hgh.chine \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
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 \<star> TfTgh.p\<^sub>1"]
Maps.comp_CLS [of Tgh_Hgh.chine TfTgh.p\<^sub>0 "Tgh_Hgh.chine \<star> TfTgh.p\<^sub>0"]
by auto
also have "... = \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk>"
using TfTgh_TfHgh.CLS_chine by simp
finally show ?thesis by blast
qed
ultimately have "Chn RHS =\<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> tuple_ABC"
by simp
also have "... = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> 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 \<a> \<i> src trg \<open>tab\<^sub>1 h\<close> \<open>tab\<^sub>0 g\<close>
using gh by unfold_locales auto
interpretation f\<^sub>0g\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg \<open>tab\<^sub>1 g\<close> \<open>tab\<^sub>0 f\<close>
using fg by unfold_locales auto
interpretation f\<^sub>0gh\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close> \<open>tab\<^sub>0 f\<close>
using fg gh Tgh.\<rho>\<sigma>.leg1_is_map
by unfold_locales auto
interpretation fg\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>tab\<^sub>1 h\<close> \<open>tab\<^sub>0 g \<star> Tfg.p\<^sub>0\<close>
using TTfgh.r\<^sub>0s\<^sub>1_is_cospan by simp
lemma src_tab_eq:
shows "(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\<theta> \<cdot> TTfgh_TfTgh.the_\<nu> =
TTfgh.tab"
proof -
have "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\<theta> \<cdot> TTfgh_TfTgh.the_\<nu> =
(\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> TTfgh.tab"
unfolding TTfgh.tab_def
using TTfgh_TfTgh.chine_is_induced_map TTfgh.tab_def TTfgh_TfTgh.\<Delta>_simps(4)
by auto
moreover have "iso (\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0)"
by (simp add: fg gh)
moreover have "inv (\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) = \<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
using fg gh by simp
ultimately show ?thesis
using TTfgh_TfTgh.\<Delta>_simps(1)
invert_side_of_triangle(1)
[of "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\<theta> \<cdot> TTfgh_TfTgh.the_\<nu>"
"\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0" TTfgh.tab]
by argo
qed
text \<open>
We need to show that the associativity isomorphism (defined in terms of tupling) coincides
with \<open>TTfgh_TfTgh.chine\<close> (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.
\<close>
lemma prj_chine:
shows "\<lbrakk>\<lbrakk>TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
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 \<open>
The required isomorphisms will each be established via \<open>T2\<close>, using the equation
\<open>src_tab_eq\<close> (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 \<open>horizontal_composite_of_arrows_of_tabulations_in_maps.prj_chine\<close>.
\<close>
define u\<^sub>f where "u\<^sub>f = g \<star> h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define w\<^sub>f where "w\<^sub>f = Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1"
define w\<^sub>f' where "w\<^sub>f' = TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine"
define \<theta>\<^sub>f
where "\<theta>\<^sub>f = (g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot> (g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot> (g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> (\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> (f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
define \<theta>\<^sub>f'
where "\<theta>\<^sub>f' = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]"
define \<beta>\<^sub>f
where "\<beta>\<^sub>f = \<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^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 \<theta>\<^sub>f: "\<guillemotleft>\<theta>\<^sub>f : tab\<^sub>0 f \<star> w\<^sub>f \<Rightarrow> u\<^sub>f\<guillemotright>"
proof (unfold \<theta>\<^sub>f_def w\<^sub>f_def u\<^sub>f_def, intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] :
tab\<^sub>0 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<guillemotright>"
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 "\<guillemotleft>f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1 :
(tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using f\<^sub>0g\<^sub>1.\<phi>_in_hom(2) Tfg.\<rho>\<sigma>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>(g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 :
(tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> ((g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using Tfg.\<rho>\<sigma>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1 :
((g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (g \<star> tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] :
(g \<star> tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> g \<star> (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>g \<star> fg\<^sub>0h\<^sub>1.\<phi> : g \<star> (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> g \<star> tab\<^sub>1 h \<star> TTfgh.p\<^sub>0\<guillemotright>"
using fg\<^sub>0h\<^sub>1.\<phi>_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>g \<star> h.tab \<star> TTfgh.p\<^sub>0 : g \<star> tab\<^sub>1 h \<star> TTfgh.p\<^sub>0 \<Rightarrow> g \<star> (h \<star> tab\<^sub>0 h) \<star> TTfgh.p\<^sub>0\<guillemotright>"
using gh fg\<^sub>0h\<^sub>1.\<phi>_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] :
g \<star> (h \<star> tab\<^sub>0 h) \<star> TTfgh.p\<^sub>0 \<Rightarrow> g \<star> h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<guillemotright>"
using gh fg\<^sub>0h\<^sub>1.\<phi>_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
qed
have \<theta>\<^sub>f': "\<guillemotleft>\<theta>\<^sub>f' : tab\<^sub>0 f \<star> w\<^sub>f' \<Rightarrow> u\<^sub>f\<guillemotright>"
proof (unfold \<theta>\<^sub>f'_def w\<^sub>f'_def u\<^sub>f_def, intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] :
tab\<^sub>0 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine \<Rightarrow> (tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using "1" "2" "3" "4" assoc'_in_hom(2) f.ide_u f.leg1_simps(3) by auto
show "\<guillemotleft>f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine :
(tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine \<Rightarrow>
((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.\<phi>_in_hom(2)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> (((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>(\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
(((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
((g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan g\<^sub>0h\<^sub>1.\<phi>_in_hom(2)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
((g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) :
((g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> g \<star> h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto
show "\<guillemotleft>g \<star> h \<star> TTfgh_TfTgh.the_\<theta> :
g \<star> h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> g \<star> h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan TTfgh_TfTgh.the_\<theta>_in_hom
by (intro hcomp_in_vhom, auto)
qed
have \<beta>\<^sub>f: "\<guillemotleft>\<beta>\<^sub>f : tab\<^sub>1 f \<star> w\<^sub>f \<Rightarrow> tab\<^sub>1 f \<star> w\<^sub>f'\<guillemotright>"
proof (unfold \<beta>\<^sub>f_def w\<^sub>f_def w\<^sub>f'_def, intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] :
tab\<^sub>1 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using TTfgh.leg1_in_hom(2) assoc'_in_hom by auto
show "\<guillemotleft>TTfgh_TfTgh.the_\<nu> :
(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using TTfgh_TfTgh.the_\<nu>_in_hom TTfgh_TfTgh.the_\<nu>_props by simp
show "\<guillemotleft>\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] :
(tab\<^sub>1 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine \<Rightarrow> tab\<^sub>1 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using 1 2 3 4 by auto
qed
have iso_\<beta>\<^sub>f: "iso \<beta>\<^sub>f"
unfolding \<beta>\<^sub>f_def
using 1 2 3 4 \<beta>\<^sub>f \<beta>\<^sub>f_def isos_compose
apply (intro isos_compose)
apply (metis TTfgh.composable TTfgh.leg1_in_hom(2) Tfg.\<rho>\<sigma>.T0.antipar(2)
Tfg.\<rho>\<sigma>.T0.ide_right Tfg.\<rho>\<sigma>.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_\<nu>_props(2) f.ide_leg1 iso_assoc by blast+
have u\<^sub>f: "ide u\<^sub>f"
using \<theta>\<^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: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>f \<Rightarrow> w\<^sub>f'\<guillemotright> \<and> \<beta>\<^sub>f = tab\<^sub>1 f \<star> \<gamma> \<and> \<theta>\<^sub>f = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>)"
proof -
have eq\<^sub>f: "f.composite_cell w\<^sub>f \<theta>\<^sub>f = f.composite_cell w\<^sub>f' \<theta>\<^sub>f' \<cdot> \<beta>\<^sub>f"
proof -
text \<open>
I don't see any alternative here to just grinding out the calculation.
The idea is to bring \<open>f.composite_cell w\<^sub>f \<theta>\<^sub>f\<close> into a form in which
\<open>src_tab_eq\<close> can be applied to eliminate \<open>\<theta>\<^sub>f\<close> in favor of \<open>\<theta>\<^sub>f'\<close>.
\<close>
have "f.composite_cell w\<^sub>f \<theta>\<^sub>f =
((f \<star> g \<star> \<a>[h, tab\<^sub>0 h, fg\<^sub>0h\<^sub>1.p\<^sub>0]) \<cdot>
(f \<star> g \<star> h.tab \<star> fg\<^sub>0h\<^sub>1.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> f\<^sub>0g\<^sub>1.p\<^sub>0, fg\<^sub>0h\<^sub>1.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, f\<^sub>0g\<^sub>1.p\<^sub>0] \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> f\<^sub>0g\<^sub>1.p\<^sub>0) \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, f\<^sub>0g\<^sub>1.p\<^sub>1, fg\<^sub>0h\<^sub>1.p\<^sub>1])) \<cdot>
\<a>[f, tab\<^sub>0 f, f\<^sub>0g\<^sub>1.p\<^sub>1 \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1] \<cdot>
(f.tab \<star> f\<^sub>0g\<^sub>1.p\<^sub>1 \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1)"
unfolding w\<^sub>f_def \<theta>\<^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 \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using comp_assoc by simp
also have "... =
(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0]) \<cdot>
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) =
f \<star> g \<star> \<a>[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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, (h \<star> tab\<^sub>0 h) \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0)) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> 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 "\<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]"]
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>1 h \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>)) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> 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 \<star> TTfgh.p\<^sub>0"]
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1]) \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> 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.\<phi>]
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot> \<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] =
\<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot> \<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] =
\<lbrace>(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]\<rbrace>"
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
\<a>'_def \<alpha>_def
by simp
also have "... =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]\<rbrace>"
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 "... = \<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> 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
\<a>'_def \<alpha>_def
by simp
finally show ?thesis by blast
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, tab\<^sub>1 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot> \<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1] =
\<a>[f, tab\<^sub>1 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> ((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> 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.\<phi>_in_hom
assoc_naturality [of f f\<^sub>0g\<^sub>1.\<phi> TTfgh.p\<^sub>1]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> \<a>[f, tab\<^sub>1 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] =
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> ((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> 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.\<phi>_in_hom
assoc_naturality [of f "g.tab \<star> Tfg.p\<^sub>0" TTfgh.p\<^sub>1]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) =
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> 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 "\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]" "\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(((\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) =
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> 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 "\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0]" "\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1)) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, (\<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>)\<rbrace>"
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
\<a>'_def \<alpha>_def
by simp
also have "... = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]\<rbrace>"
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 "... = \<a>[f \<star> g, tab\<^sub>0 g \<star> 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
\<a>'_def \<alpha>_def
by simp
finally show ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) =
((f \<star> g) \<star> \<a>[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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
TTfgh.tab \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using TTfgh.tab_def Tfg.\<rho>\<sigma>.tab_def by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<cdot>
(g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> Tgh.p\<^sub>0) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1)
\<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1)
\<star> TTfgh_TfTgh.chine) \<cdot>
TTfgh_TfTgh.the_\<nu>) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using src_tab_eq TfTgh.tab_def Tgh.\<rho>\<sigma>.tab_def comp_assoc by simp
text \<open>Now we have to make this look like \<open>f.composite_cell w\<^sub>f' \<theta>\<^sub>f' \<cdot> \<beta>\<^sub>f\<close>.\<close>
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>)) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1)
\<star> TTfgh_TfTgh.chine) \<cdot>
TTfgh_TfTgh.the_\<nu>) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<cdot>
(g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> Tgh.p\<^sub>0) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1)
\<star> TfTgh.p\<^sub>0 =
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0)"
using fg gh whisker_right whisker_left by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1)
\<star> TTfgh_TfTgh.chine =
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine)"
(* using fg gh whisker_right [of TTfgh_TfTgh.chine] by auto (* 2 min *) *)
proof -
have "arr (\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using fg gh
by (intro seqI' comp_in_homI) auto
moreover
have "arr ((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr (\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> 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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
(((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
(f.tab \<star> TfTgh.p\<^sub>1) \<star> 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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
(((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> 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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine]) \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine]) \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) =
f.tab \<star> TfTgh.p\<^sub>1 \<star> 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 "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
proof -
have "(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] =
\<lbrace>(\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<a>'_def \<alpha>_def by simp
also have "... = \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>])\<rbrace>"
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 "... = \<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
using \<a>'_def \<alpha>_def by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> 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.\<phi>, g\<^sub>0h\<^sub>1.\<phi>, h.tab, with associativities that
* do not respect their domain and codomain.
*
* I also tried to avoid distributing the "f \<star>" 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 \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> 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.\<phi> TTfgh_TfTgh.chine] comp_assoc
by simp
also have "... =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, ((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, ((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> 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 \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, ((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> 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 "\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> 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 \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
(((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> 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 \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
(* OK, we can perhaps shortcut the last few steps... *)
have "((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<lbrace>((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>])
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
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
\<a>'_def \<alpha>_def
by simp
also have "... =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>])
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)\<rbrace>"
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 "... =
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> 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
\<a>'_def \<alpha>_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 "... =
(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>)) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) =
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^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 \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(((f \<star> g) \<star> h) \<star> TTfgh_TfTgh.the_\<theta>)) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
proof -
have "(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> ((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) =
\<a>\<^sup>-\<^sup>1[f, g, h] \<star> TTfgh_TfTgh.the_\<theta>"
using fg gh comp_arr_dom comp_cod_arr
interchange [of "\<a>\<^sup>-\<^sup>1[f, g, h]" "f \<star> g \<star> h"
"tab\<^sub>0 h \<star> TTfgh.p\<^sub>0" TTfgh_TfTgh.the_\<theta>]
by simp
also have "... = (((f \<star> g) \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh comp_arr_dom comp_cod_arr
interchange [of "(f \<star> g) \<star> h" "\<a>\<^sup>-\<^sup>1[f, g, h]" TTfgh_TfTgh.the_\<theta>
"((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine"]
by simp
finally show ?thesis
using comp_assoc by presburger
qed
also have "... =
(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> h \<star> TTfgh_TfTgh.the_\<theta>)) \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh assoc_naturality [of "f \<star> g" h TTfgh_TfTgh.the_\<theta>] comp_assoc
by simp
also have "... =
(f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh assoc_naturality [of f g "h \<star> TTfgh_TfTgh.the_\<theta>"] comp_assoc
by simp
finally show ?thesis
using comp_assoc by presburger
qed
also have "... =
((f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(f \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) =
f \<star> can
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>,
\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>,
((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>\<^bold>]
\<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, (\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>,
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star>
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)\<rbrace>"
using \<a>'_def \<alpha>_def by simp
also have "... =
can (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>))
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>))"
using fg gh
apply (unfold can_def)
apply (intro E.eval_eqI)
by simp_all
also have "... =
f \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
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 \<star> (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(f \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) =
f \<star> (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]"
proof -
have 1: "arr ((g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^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>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^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' \<theta>\<^sub>f' \<cdot> \<beta>\<^sub>f"
unfolding w\<^sub>f'_def \<theta>\<^sub>f'_def \<beta>\<^sub>f_def
using comp_assoc by presburger
finally show ?thesis by blast
qed
show ?thesis
using w\<^sub>f w\<^sub>f' \<theta>\<^sub>f \<theta>\<^sub>f' \<beta>\<^sub>f f.T2 [of w\<^sub>f w\<^sub>f' \<theta>\<^sub>f u\<^sub>f \<theta>\<^sub>f' \<beta>\<^sub>f] eq\<^sub>f by fast
qed
obtain \<gamma>\<^sub>f where \<gamma>\<^sub>f: "\<guillemotleft>\<gamma>\<^sub>f : w\<^sub>f \<Rightarrow> w\<^sub>f'\<guillemotright> \<and> \<beta>\<^sub>f = tab\<^sub>1 f \<star> \<gamma>\<^sub>f \<and> \<theta>\<^sub>f = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>\<^sub>f)"
using 5 by auto
show "\<lbrakk>\<lbrakk>TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "iso \<gamma>\<^sub>f"
using \<gamma>\<^sub>f BS3 w\<^sub>f_is_map w\<^sub>f'_is_map by blast
hence "isomorphic w\<^sub>f w\<^sub>f'"
using \<gamma>\<^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 \<open>
On to the next equation:
\[
\<open>\<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>\<close>.
\]
We have to make use of the equation \<open>\<theta>\<^sub>f = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>\<^sub>f)\<close> in this part,
similarly to how the equation \<open>src_tab_eq\<close> was used to replace
\<open>TTfgh.tab\<close> in the first part.
\<close>
define u\<^sub>g where "u\<^sub>g = h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define w\<^sub>g where "w\<^sub>g = Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1"
define w\<^sub>g' where "w\<^sub>g' = Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
define \<theta>\<^sub>g
where "\<theta>\<^sub>g = \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
define \<theta>\<^sub>g'
where "\<theta>\<^sub>g' = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
define \<beta>\<^sub>g
where "\<beta>\<^sub>g = \<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot> \<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^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 \<theta>\<^sub>g: "\<guillemotleft>\<theta>\<^sub>g : tab\<^sub>0 g \<star> w\<^sub>g \<Rightarrow> u\<^sub>g\<guillemotright>"
using w\<^sub>g_def u\<^sub>g_def \<theta>\<^sub>g_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\<phi>_in_hom by auto
have \<theta>\<^sub>g': "\<guillemotleft>\<theta>\<^sub>g' : tab\<^sub>0 g \<star> w\<^sub>g' \<Rightarrow> u\<^sub>g\<guillemotright>"
unfolding w\<^sub>g'_def u\<^sub>g_def \<theta>\<^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 \<beta>\<^sub>g: "\<guillemotleft>\<beta>\<^sub>g : tab\<^sub>1 g \<star> w\<^sub>g \<Rightarrow> tab\<^sub>1 g \<star> w\<^sub>g'\<guillemotright>"
proof (unfold \<beta>\<^sub>g_def w\<^sub>g_def, intro comp_in_homI)
(* auto can solve this, but it's too slow *)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] :
tab\<^sub>1 g \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1 :
(tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\<phi>_in_hom f\<^sub>0g\<^sub>1.\<phi>_uniqueness(2)
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] :
(tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \<gamma>\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto
show "\<guillemotleft>tab\<^sub>0 f \<star> \<gamma>\<^sub>f : tab\<^sub>0 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \<gamma>\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] :
tab\<^sub>0 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine \<Rightarrow> (tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine\<guillemotright>"
by auto
show "\<guillemotleft>f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine :
(tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.\<phi>_in_hom
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] :
((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> (tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
by auto
show "\<guillemotleft>\<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] :
(tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine \<Rightarrow> tab\<^sub>1 g \<star> w\<^sub>g'\<guillemotright>"
using w\<^sub>g'_def by auto
qed
have eq\<^sub>g: "g.composite_cell w\<^sub>g \<theta>\<^sub>g = g.composite_cell w\<^sub>g' \<theta>\<^sub>g' \<cdot> \<beta>\<^sub>g"
proof -
have "g.composite_cell w\<^sub>g \<theta>\<^sub>g =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot>
(h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
fg\<^sub>0h\<^sub>1.\<phi> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
unfolding w\<^sub>g_def \<theta>\<^sub>g_def by simp
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
((g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1]) \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> 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 \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]))) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
proof -
have "(\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) =
g \<star> \<a>\<^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 \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1]) \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
using comp_assoc by presburger
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> 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 "\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1"
"(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1])"
"\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" "\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1]"]
by simp
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^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 \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
proof -
have "(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] =
((f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^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.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
((tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^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.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^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.\<phi>_uniqueness comp_cod_arr by simp
also have "... = ((tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
proof -
have "(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot> (inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) =
f\<^sub>0g\<^sub>1.\<phi> \<cdot> inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1"
using f\<^sub>0g\<^sub>1.\<phi>_uniqueness whisker_right by simp
also have "... = (tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1"
using f\<^sub>0g\<^sub>1.\<phi>_uniqueness comp_arr_inv' by simp
finally show ?thesis by simp
qed
also have "... = \<a>\<^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.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = \<a>\<^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 "... = \<theta>\<^sub>f \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
unfolding \<theta>\<^sub>f_def using comp_assoc by presburger
also have "... = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using \<gamma>\<^sub>f comp_assoc by simp
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
unfolding \<theta>\<^sub>f'_def using comp_assoc by presburger
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
proof -
have "(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) =
f\<^sub>0gh\<^sub>1.\<phi> \<star> 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 \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<beta>\<^sub>g"
unfolding \<beta>\<^sub>g_def using comp_assoc by presburger
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<beta>\<^sub>g"
proof -
have "(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of "(g.tab \<star> Tgh.p\<^sub>1)" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
proof -
have "((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> 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 \<star> TTfgh_TfTgh.chine"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<beta>\<^sub>g"
proof -
have "(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) =
g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
using comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc g\<^sub>0h\<^sub>1.\<phi>_in_hom by simp
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<beta>\<^sub>g"
proof -
have "(\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) =
(\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> (tab\<^sub>0 g \<star> Tgh.p\<^sub>1)) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> 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 "... = (\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> 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 \<star> (tab\<^sub>0 g \<star> Tgh.p\<^sub>1)) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> 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 "... = (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> 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 \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<beta>\<^sub>g"
using comp_assoc by presburger
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
proof -
have "(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
proof -
have "(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
using comp_assoc by presburger
also have "... = ((((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
proof -
have "(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of "g \<star> g\<^sub>0h\<^sub>1.\<phi>" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
proof -
have "(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of "g \<star> h.tab \<star> Tgh.p\<^sub>0" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
proof -
have "((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of g g\<^sub>0h\<^sub>1.\<phi> "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
proof -
have "((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of g "h.tab \<star> Tgh.p\<^sub>0" "TfTgh.p\<^sub>0 \<star> 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 \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(g \<star> can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine])) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
proof -
have "can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
g \<star> can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
can (((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \<a>'_def \<alpha>_def by simp
also have "... = can (((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
unfolding can_def
using gh
apply (intro E.eval_eqI) by simp_all
finally show ?thesis by blast
qed
moreover
have "\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
can ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \<a>'_def \<alpha>_def by simp
also have "... = can ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
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 "\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
proof -
have "\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \<a>'_def \<alpha>_def by simp
also have "... = \<lbrace>\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
apply (intro E.eval_eqI) by simp_all
also have "... = g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \<a>'_def \<alpha>_def by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using comp_assoc by presburger
qed
also have "... = (g \<star>
(h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
using gh whisker_left by auto (* 11 sec *)
also have "... = g.composite_cell w\<^sub>g' \<theta>\<^sub>g' \<cdot> \<beta>\<^sub>g"
unfolding w\<^sub>g'_def \<theta>\<^sub>g'_def
using comp_assoc by presburger
finally show ?thesis by blast
qed
have 6: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>g \<Rightarrow> w\<^sub>g'\<guillemotright> \<and> \<beta>\<^sub>g = tab\<^sub>1 g \<star> \<gamma> \<and> \<theta>\<^sub>g = \<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>)"
using w\<^sub>g w\<^sub>g' \<theta>\<^sub>g \<theta>\<^sub>g' \<beta>\<^sub>g eq\<^sub>g g.T2 [of w\<^sub>g w\<^sub>g' \<theta>\<^sub>g u\<^sub>g \<theta>\<^sub>g' \<beta>\<^sub>g] by blast
obtain \<gamma>\<^sub>g where \<gamma>\<^sub>g: "\<guillemotleft>\<gamma>\<^sub>g : w\<^sub>g \<Rightarrow> w\<^sub>g'\<guillemotright> \<and> \<beta>\<^sub>g = tab\<^sub>1 g \<star> \<gamma>\<^sub>g \<and> \<theta>\<^sub>g = \<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g)"
using 6 by auto
show "\<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "iso \<gamma>\<^sub>g"
using \<gamma>\<^sub>g BS3 w\<^sub>g_is_map w\<^sub>g'_is_map by blast
hence "isomorphic w\<^sub>g w\<^sub>g'"
using \<gamma>\<^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 \<open>Now the last equation: similar, but somewhat simpler.\<close>
define u\<^sub>h where "u\<^sub>h = tab\<^sub>0 h \<star> 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 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
define \<theta>\<^sub>h
where "\<theta>\<^sub>h = tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define \<theta>\<^sub>h'
where "\<theta>\<^sub>h' = TTfgh_TfTgh.the_\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
define \<beta>\<^sub>h
where "\<beta>\<^sub>h = \<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
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 \<theta>\<^sub>h: "\<guillemotleft>\<theta>\<^sub>h : tab\<^sub>0 h \<star> w\<^sub>h \<Rightarrow> u\<^sub>h\<guillemotright>"
unfolding \<theta>\<^sub>h_def w\<^sub>h_def u\<^sub>h_def by auto
have \<theta>\<^sub>h': "\<guillemotleft>\<theta>\<^sub>h' : tab\<^sub>0 h \<star> w\<^sub>h' \<Rightarrow> u\<^sub>h\<guillemotright>"
unfolding \<theta>\<^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 \<beta>\<^sub>h: "\<guillemotleft>\<beta>\<^sub>h : tab\<^sub>1 h \<star> w\<^sub>h \<Rightarrow> tab\<^sub>1 h \<star> w\<^sub>h'\<guillemotright>"
proof (unfold \<beta>\<^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 "\<guillemotleft>inv fg\<^sub>0h\<^sub>1.\<phi> : tab\<^sub>1 h \<star> TTfgh.p\<^sub>0 \<Rightarrow> (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.\<phi>_uniqueness by blast
show "\<guillemotleft>\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] :
(tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 g \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>tab\<^sub>0 g \<star> \<gamma>\<^sub>g :
tab\<^sub>0 g \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 g \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using \<gamma>\<^sub>g w\<^sub>g_def w\<^sub>g'_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] :
tab\<^sub>0 g \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine
\<Rightarrow> (tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine :
(tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine
\<Rightarrow> (tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by force
show "\<guillemotleft>\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] :
(tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine
\<Rightarrow> tab\<^sub>1 h \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
qed
have eq\<^sub>h: "h.composite_cell w\<^sub>h \<theta>\<^sub>h = h.composite_cell w\<^sub>h' \<theta>\<^sub>h' \<cdot> \<beta>\<^sub>h"
proof -
text \<open>
Once again, the strategy is to form the subexpression
\[
\<open>\<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]\<close>
\]
which is equal to \<open>\<theta>\<^sub>g\<close>, so that we can make use of the equation \<open>\<theta>\<^sub>g = \<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g)\<close>.
\<close>
have "h.composite_cell w\<^sub>h \<theta>\<^sub>h =
(h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0)"
unfolding w\<^sub>h_def \<theta>\<^sub>h_def by simp
also have "... = \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0)"
proof -
have "(h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] = \<a>[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 "... = (\<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
fg\<^sub>0h\<^sub>1.\<phi> \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "(h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> (\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi> =
(h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> ((tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc_assoc' by simp
also have "... = (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\<phi>_uniqueness comp_cod_arr by simp
also have "... = (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> (tab\<^sub>1 h \<star> TTfgh.p\<^sub>0)"
using comp_arr_inv' fg\<^sub>0h\<^sub>1.\<phi>_uniqueness by simp
also have "... = h.tab \<star> TTfgh.p\<^sub>0"
using comp_arr_dom fg\<^sub>0h\<^sub>1.p\<^sub>0_simps by simp
finally have "(h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> (\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi> =
h.tab \<star> TTfgh.p\<^sub>0"
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<theta>\<^sub>g \<cdot> \<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
unfolding \<theta>\<^sub>g_def by simp
also have "... = (\<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g)) \<cdot> \<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
using \<gamma>\<^sub>g by simp
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
unfolding \<theta>\<^sub>g'_def
using comp_assoc by presburger
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) =
(h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
using comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using comp_assoc by presburger
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using assoc_naturality [of h.tab Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"] comp_assoc
by simp
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "(\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) =
h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
using comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine])) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using comp_assoc by presburger
also have "... = ((h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(h \<star> can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>))) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) =
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<a>'_def \<alpha>_def by simp
also have "... = can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
can (((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace> =
can (((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
unfolding can_def
apply (intro E.eval_eqI) by simp_all
thus ?thesis by simp
qed
also have "... = can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
by simp
also have "... = h \<star> can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
using whisker_can_left_0 by simp
finally have "can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) =
h \<star> can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta> \<cdot>
can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using whisker_left [of h] comp_assoc by simp
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
unfolding can_def
apply (intro E.eval_eqI) by auto
also have "... = \<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
using \<a>'_def \<alpha>_def by simp
finally have "can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) =
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
by simp
thus ?thesis by simp
qed
also have "... = h.composite_cell w\<^sub>h' \<theta>\<^sub>h' \<cdot> \<beta>\<^sub>h"
unfolding w\<^sub>h'_def \<theta>\<^sub>h'_def \<beta>\<^sub>h_def
using comp_assoc by presburger
finally show ?thesis by simp
qed
have 7: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>h \<Rightarrow> w\<^sub>h'\<guillemotright> \<and> \<beta>\<^sub>h = tab\<^sub>1 h \<star> \<gamma> \<and> \<theta>\<^sub>h = \<theta>\<^sub>h' \<cdot> (tab\<^sub>0 h \<star> \<gamma>)"
using w\<^sub>h w\<^sub>h' \<theta>\<^sub>h \<theta>\<^sub>h' \<beta>\<^sub>h eq\<^sub>h h.T2 [of w\<^sub>h w\<^sub>h' \<theta>\<^sub>h u\<^sub>h \<theta>\<^sub>h' \<beta>\<^sub>h] by blast
obtain \<gamma>\<^sub>h where \<gamma>\<^sub>h: "\<guillemotleft>\<gamma>\<^sub>h : w\<^sub>h \<Rightarrow> w\<^sub>h'\<guillemotright> \<and> \<beta>\<^sub>h = tab\<^sub>1 h \<star> \<gamma>\<^sub>h \<and> \<theta>\<^sub>h = \<theta>\<^sub>h' \<cdot> (tab\<^sub>0 h \<star> \<gamma>\<^sub>h)"
using 7 by auto
show "\<lbrakk>\<lbrakk>Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "iso \<gamma>\<^sub>h"
using \<gamma>\<^sub>h BS3 w\<^sub>h_is_map w\<^sub>h'_is_map by blast
hence "isomorphic w\<^sub>h w\<^sub>h'"
using \<gamma>\<^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 \<open>
Finally, we can show that @{term TTfgh_TfTgh.chine} is given by tupling.
\<close>
lemma CLS_chine:
shows "\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = tuple_ABC"
proof -
have "tuple_ABC = SPN_fgh.chine_assoc"
using SPN_fgh.chine_assoc_def by simp
also have "... = \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof (intro Maps.arr_eqI)
show "Maps.arr SPN_fgh.chine_assoc"
using SPN_fgh.chine_assoc_in_hom by auto
show "Maps.arr \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using Maps.CLS_in_hom TTfgh_TfTgh.is_map by blast
show "Maps.Dom SPN_fgh.chine_assoc = Maps.Dom \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 \<star> 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 \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 = \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof (intro Maps.prj_joint_monic
[of SPN_fgh.\<mu>.leg0 "SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1"
tuple_ABC "\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"])
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using SPN_fgh.\<nu>\<pi>.dom.is_span SPN_fgh.\<nu>\<pi>.leg1_composite SPN_fgh.cospan_\<mu>\<nu>
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 \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof
show "Maps.in_hom \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>src TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 \<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> SPN_fgh.\<mu>.apex"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using SPN_fgh.prj_in_hom(4) by blast
show "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
Maps.pbdom SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
proof -
have "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = Maps.MkIde (src TfTgh.p\<^sub>0)"
by simp
also have "... = Maps.pbdom SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger
finally show ?thesis by blast
qed
show "SPN_fgh.\<mu>.apex = Maps.dom SPN_fgh.\<mu>.leg0"
using SPN_f.dom.apex_def by blast
qed
qed
have 2: "Maps.commutative_square SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1
SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
proof
show "Maps.cospan SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1"
using SPN_fgh.\<nu>\<pi>.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.\<nu>.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.\<nu>.leg0 \<odot> SPN_fgh.Prj\<^sub>0\<^sub>1 = SPN_fgh.\<pi>.leg1 \<odot> SPN_fgh.Prj\<^sub>0"
by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc
Maps.pullback_commutes' SPN_fgh.\<mu>\<nu>.dom.leg_simps(1)
SPN_fgh.\<mu>\<nu>.leg0_composite SPN_fgh.cospan_\<nu>\<pi>)
qed
have 1: "Maps.commutative_square
SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.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.\<mu>.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.\<mu>.leg0 \<odot> SPN_fgh.Prj\<^sub>1\<^sub>1 =
(SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_BC"
using 2 fg gh Maps.comp_assoc csq(2)
Maps.prj_tuple [of SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0]
by blast
qed
show "SPN_fgh.Prj\<^sub>1 \<odot> tuple_ABC = SPN_fgh.Prj\<^sub>1 \<odot> Maps.CLS TTfgh_TfTgh.chine"
proof -
have "SPN_fgh.Prj\<^sub>1 \<odot> tuple_ABC = SPN_fgh.Prj\<^sub>1\<^sub>1"
using csq(2) by simp
also have "... = \<lbrakk>\<lbrakk>Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using prj_char by simp
also have "... = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using prj_chine(1) by simp
also have "... = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 \<star> TTfgh_TfTgh.chine \<cong> TfTgh.p\<^sub>1 \<star> 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 \<star> TTfgh_TfTgh.chine"]
by simp
qed
also have "... = SPN_fgh.Prj\<^sub>1 \<odot> Maps.CLS TTfgh_TfTgh.chine"
using prj_char by simp
finally show ?thesis by blast
qed
show
"Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_ABC =
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have
"Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_ABC =
tuple_BC"
using csq(2)
Maps.prj_tuple [of SPN_fgh.\<mu>.leg0 "SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1"
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC]
by simp
also have "... =
Maps.comp
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1))
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof (intro Maps.prj_joint_monic
[of SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 tuple_BC
"Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"])
show "Maps.cospan SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1"
using SPN_fgh.\<nu>\<pi>.legs_form_cospan(1) by simp
show "Maps.seq SPN_fgh.\<nu>\<pi>.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.\<nu>\<pi>.prj\<^sub>1 (Maps.MkIde (src Tgh.p\<^sub>0)) SPN_fgh.\<nu>.apex"
proof -
have "Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 = Maps.MkIde (src Tgh.p\<^sub>0)"
using fg gh Maps.pbdom_def
by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.\<nu>\<pi>.are_identities(2)
SPN_fgh.\<nu>\<pi>.composable Span.chine_hcomp_ide_ide Tgh.chine_hcomp_SPN_SPN
g.is_ide)
thus ?thesis
using SPN_fgh.\<nu>\<pi>.prj_in_hom(1) by simp
qed
qed
show "Maps.seq SPN_fgh.\<nu>\<pi>.prj\<^sub>1
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>)"
proof
show "Maps.in_hom SPN_fgh.\<nu>\<pi>.prj\<^sub>1
(Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1) SPN_fgh.\<nu>.apex"
using SPN_fgh.\<nu>\<pi>.prj_in_hom(1) by simp
show "Maps.in_hom
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>src TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
(Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1)"
proof
show "Maps.in_hom \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>src TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1))
\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
(Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1)"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using SPN_fgh.prj_in_hom(4) by blast
show "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
Maps.pbdom SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
proof -
have "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = Maps.MkIde (src TfTgh.p\<^sub>0)"
by simp
also have "... = Maps.pbdom SPN_fgh.\<mu>.leg0
(SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.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.\<nu>.leg0 SPN_fgh.\<pi>.leg1 =
Maps.dom (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using fg gh Maps.pbdom_def SPN_fgh.\<nu>\<pi>.apex_composite
SPN_fgh.\<nu>\<pi>.dom.apex_def SPN_fgh.\<nu>\<pi>.dom.is_span
SPN_fgh.\<nu>\<pi>.leg1_composite
by presburger
qed
qed
qed
show "SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot> tuple_BC =
SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot> tuple_BC = SPN_fgh.Prj\<^sub>0"
using csq(1) by simp
also have "... = SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>Tgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.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 \<star> Tgh.p\<^sub>1"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive
Maps.comp_CLS
[of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive
Maps.comp_CLS [of Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
"Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
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.\<nu>\<pi>.prj\<^sub>1 \<odot> tuple_BC =
SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot> tuple_BC = SPN_fgh.Prj\<^sub>0\<^sub>1"
using csq(1) by simp
also have "... = SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>Tgh.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.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 \<star> Tgh.p\<^sub>1"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive
Maps.comp_CLS
[of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive
Maps.comp_CLS [of Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
"Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
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 \<open>
At long last, we can show associativity coherence for \<open>SPN\<close>.
\<close>
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 = \<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
using Chn_LHS_eq by simp
also have "... = \<lbrakk>\<lbrakk>HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "\<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
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 \<star> TTfgh_THfgh.chine"]
by simp
moreover
have "\<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "ide (HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine)"
proof -
have "ide (THfgh_HHfgh.chine \<star> 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 \<star> TTfgh_THfgh.chine)"
using fg gh HHfgh_HfHgh.chine_in_hom \<alpha>_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 \<star> TTfgh_THfgh.chine"
"HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine"]
by auto
qed
ultimately show ?thesis by argo
qed
also have "... = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
interpret A: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfTgh.tab \<open>(tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfTgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>\<a>[f, g, h]\<close> \<open>f \<star> g \<star> h\<close>
..
interpret B: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> HfHgh.tab \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>\<a>[f, g, h]\<close> \<open>f \<star> g \<star> h\<close>
using fg gh by unfold_locales auto
interpret C: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>(f \<star> g) \<star> h\<close> HHfgh.tab \<open>tab\<^sub>0 ((f \<star> g) \<star> h)\<close> \<open>tab\<^sub>1 ((f \<star> g) \<star> h)\<close>
\<open>f \<star> g \<star> h\<close> HfHgh.tab \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>(f \<star> g) \<star> h\<close> \<open>\<a>[f, g, h]\<close>
using fg gh by unfold_locales auto
interpret D: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> HfHgh.tab \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>(f \<star> g) \<star> h\<close> \<open>\<a>[f, g, h]\<close>
using fg gh by unfold_locales auto
have "HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine \<cong> D.chine"
proof -
have "D.chine \<cong> D.\<pi>.chine \<star> TTfgh_THfgh.chine"
using D.chine_char by simp
also have "... \<cong> C.chine \<star> TTfgh_THfgh.chine"
using fg gh comp_arr_dom isomorphic_reflexive by simp
also have "... \<cong> (C.\<pi>.chine \<star> THfgh_HHfgh.chine) \<star> TTfgh_THfgh.chine"
using C.chine_char hcomp_isomorphic_ide by simp
also have "... \<cong> (HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine) \<star> TTfgh_THfgh.chine"
proof -
have "C.\<pi>.chine = HHfgh_HfHgh.chine"
using fg gh comp_arr_dom comp_cod_arr \<alpha>_def by simp
hence "isomorphic C.\<pi>.chine HHfgh_HfHgh.chine"
using isomorphic_reflexive by simp
thus ?thesis
using hcomp_isomorphic_ide by simp
qed
also have "... \<cong> HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine"
proof -
have "ide HHfgh_HfHgh.chine \<and> ide THfgh_HHfgh.chine \<and> ide TTfgh_THfgh.chine"
by simp
moreover have "src HHfgh_HfHgh.chine = trg THfgh_HHfgh.chine \<and>
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 \<alpha>_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 "... \<cong> 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 "... \<cong> TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
proof -
have "B.chine \<cong> TfHgh_HfHgh.chine \<star> B.\<mu>.chine"
using B.chine_char by simp
also have "... \<cong> TfHgh_HfHgh.chine \<star> A.chine"
using fg gh comp_cod_arr isomorphic_reflexive by simp
also have "... \<cong> TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
using A.chine_char hcomp_ide_isomorphic by simp
finally show ?thesis by blast
qed
finally have "HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine \<cong>
TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
by blast
thus ?thesis
using fg gh Maps.CLS_eqI isomorphic_implies_hpar(1) by blast
qed
also have "... = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
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 \<star> TTfgh_TfTgh.chine"]
Maps.comp_CLS
[of TfHgh_HfHgh.chine "TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
"TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> 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 \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 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 \<open>\<lambda>\<mu>\<nu>. Span.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
..
interpretation SPNoH: composite_functor VV.comp V Span.vcomp
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> SPN
..
interpretation \<Phi>: transformation_by_components VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<open>\<lambda>rs. CMP (fst rs) (snd rs)\<close>
using compositor_is_natural_transformation by simp
interpretation \<Phi>: natural_isomorphism VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<Phi>.map
using compositor_is_natural_isomorphism by simp
abbreviation \<Phi>
where "\<Phi> \<equiv> \<Phi>.map"
interpretation SPN: pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>
proof
show "\<And>f g h. \<lbrakk> ide f; ide g; ide h; src f = trg g; src g = trg h \<rbrakk> \<Longrightarrow>
SPN \<a>[f, g, h] \<bullet> \<Phi> (f \<star> g, h) \<bullet> (\<Phi> (f, g) \<circ> SPN h) =
\<Phi> (f, g \<star> h) \<bullet> (SPN f \<circ> \<Phi> (g, h)) \<bullet> 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 \<a> \<i> 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 \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>"
..
interpretation SPN: equivalence_pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>
proof
show "\<And>\<mu> \<mu>'. \<lbrakk>par \<mu> \<mu>'; SPN \<mu> = SPN \<mu>'\<rbrakk> \<Longrightarrow> \<mu> = \<mu>'"
proof -
fix \<mu> \<mu>'
assume par: "par \<mu> \<mu>'"
assume eq: "SPN \<mu> = SPN \<mu>'"
interpret dom_\<mu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close>
using par apply unfold_locales by auto
interpret cod_\<mu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>cod \<mu>\<close>
using par apply unfold_locales by auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<open>tab_of_ide (dom \<mu>)\<close> \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<open>tab_of_ide (cod \<mu>)\<close> \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>
using par apply unfold_locales by auto
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>
using par apply unfold_locales by auto
interpret \<mu>': arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<open>tab_of_ide (dom \<mu>)\<close> \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<open>tab_of_ide (cod \<mu>)\<close> \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>'
using par apply unfold_locales by auto
interpret \<mu>': arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>'
using par apply unfold_locales by auto
have "\<mu>.chine \<cong> \<mu>'.chine"
using par eq SPN_def spn_def Maps.CLS_eqI \<mu>.is_ide by auto
hence "\<mu>.\<Delta> = \<mu>'.\<Delta>"
using \<mu>.\<Delta>_naturality \<mu>'.\<Delta>_naturality
by (metis \<mu>.\<Delta>_simps(4) \<mu>'.\<Delta>_simps(4) \<mu>.chine_is_induced_map \<mu>'.chine_is_induced_map
\<mu>.induced_map_preserved_by_iso)
thus "\<mu> = \<mu>'"
using par \<mu>.\<mu>_in_terms_of_\<Delta> \<mu>'.\<mu>_in_terms_of_\<Delta> by metis
qed
show "\<And>a'. Span.obj a' \<Longrightarrow> \<exists>a. obj a \<and> 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 \<lbrakk>\<lbrakk>tab\<^sub>0 (dom (Maps.Dom (Chn a')))\<rbrakk>\<rbrakk>"
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
\<open>Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\<close> obj_def)
by (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src
\<open>Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\<close> obj_def)
qed
thus ?thesis
using Span.equivalent_objects_reflexive
by (simp add: a')
qed
ultimately show "\<exists>a. obj a \<and> Span.equivalent_objects (SPN.map\<^sub>0 a) a'"
by auto
qed
show "\<And>a b g. \<lbrakk>obj a; obj b; Span.in_hhom g (SPN.map\<^sub>0 a) (SPN.map\<^sub>0 b); Span.ide g\<rbrakk>
\<Longrightarrow> \<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> ide f \<and> 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 "\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> ide f \<and> 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 \<a> \<i> src trg \<open>Maps.REP g.leg0\<close>
using Maps.REP_in_Map [of g.leg0]
by unfold_locales auto
have 0: "\<guillemotleft>Maps.REP g.leg0 : src (Maps.REP g.apex) \<rightarrow> Maps.Cod g.leg0\<guillemotright>"
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: "\<guillemotleft>Maps.REP g.leg1 : src (Maps.REP g.apex) \<rightarrow> Maps.Cod g.leg1\<guillemotright>"
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 \<star> (Maps.REP g.leg0)\<^sup>*"
have f_in_hhom: "\<guillemotleft>?f : a \<rightarrow> b\<guillemotright>"
proof
show "\<guillemotleft>Maps.REP g.leg1 : src (Maps.REP g.apex) \<rightarrow> b\<guillemotright>"
proof -
have "\<guillemotleft>Maps.REP g.leg1 : src (Maps.REP g.apex) \<rightarrow> Maps.Cod g.leg1\<guillemotright>"
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 \<lbrakk>\<lbrakk>tab\<^sub>0 b\<rbrakk>\<rbrakk>))"
using b arr_b SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>trg (tab\<^sub>0 b)\<rbrakk>\<rbrakk>)"
using b Maps.CLS_in_hom [of "tab\<^sub>0 b"] by force
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>b\<rbrakk>\<rbrakk>)"
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 "\<guillemotleft>(Maps.REP g.leg0)\<^sup>* : a \<rightarrow> src (Maps.REP g.apex)\<guillemotright>"
proof -
have "\<guillemotleft>Maps.REP g.leg0 : src (Maps.REP g.apex) \<rightarrow> a\<guillemotright>"
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 \<lbrakk>\<lbrakk>tab\<^sub>0 a\<rbrakk>\<rbrakk>))"
using a arr_a SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>trg (tab\<^sub>0 a)\<rbrakk>\<rbrakk>)"
using a Maps.CLS_in_hom [of "tab\<^sub>0 a"] by force
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>a\<rbrakk>\<rbrakk>)"
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 = \<lparr>Chn = \<lbrakk>\<lbrakk>spn ?f\<rbrakk>\<rbrakk>,
Dom = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 ?f\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 ?f\<rbrakk>\<rbrakk>\<rparr>,
Cod = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 ?f\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 ?f\<rbrakk>\<rbrakk>\<rparr>\<rparr>"
using calculation(1) SPN_def [of ?f] REP_leg0.antipar by auto
text \<open>
We need an invertible arrow of spans from \<open>SPN f\<close> to \<open>g\<close>.
There exists a tabulation \<open>(REP g.leg0, \<rho>, REP g.leg1)\<close> of \<open>f\<close>.
There is also a tabulation \<open>(tab\<^sub>0 f, \<rho>', tab\<^sub>1 f) of f\<close>.
As these are tabulations of the same arrow, they are equivalent.
This yields an equivalence map which is an arrow of spans from
\<open>(tab\<^sub>0 f, tab\<^sub>1 f)\<close> to \<open>(REP g.leg0, \<rho>, REP g.leg1)\<close>.
Its isomorphism class is an invertible arrow of spans in maps
from \<open>(CLS (tab\<^sub>0 f), CLS (tab\<^sub>1 f))\<close> to \<open>(g.leg0, g.leg1)\<close>.
\<close>
interpret f: identity_in_bicategory_of_spans V H \<a> \<i> src trg ?f
using ide_f apply unfold_locales by auto
interpret f: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
?f f.tab \<open>tab\<^sub>0 ?f\<close> \<open>tab\<^sub>1 ?f\<close> ?f f.tab \<open>tab\<^sub>0 ?f\<close> \<open>tab\<^sub>1 ?f\<close> ?f
using f.is_arrow_of_tabulations_in_maps by simp
interpret g: span_of_maps V H \<a> \<i> src trg \<open>Maps.REP g.leg0\<close> \<open>Maps.REP g.leg1\<close>
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 "\<exists>\<rho>. tabulation (\<cdot>) (\<star>) \<a> \<i> src trg ?f \<rho> (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 \<a> \<i> src trg ?f
(REP_leg0.trnr\<^sub>\<eta> (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: "\<exists>w w' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'.
equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src (tab\<^sub>0 ?f) \<rightarrow> src (Maps.REP g.leg0)\<guillemotright> \<and>
\<guillemotleft>w' : src (Maps.REP g.leg0) \<rightarrow> src (tab\<^sub>0 ?f)\<guillemotright> \<and>
\<guillemotleft>\<theta> : Maps.REP g.leg0 \<star> w \<Rightarrow> tab\<^sub>0 ?f\<guillemotright> \<and>
\<guillemotleft>\<nu> : tab\<^sub>1 ?f \<Rightarrow> Maps.REP g.leg1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
f.tab = (?f \<star> \<theta>) \<cdot> \<a>[?f, Maps.REP g.leg0, w] \<cdot>
(REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : tab\<^sub>0 ?f \<star> w' \<Rightarrow> Maps.REP g.leg0\<guillemotright> \<and>
\<guillemotleft>\<nu>' : Maps.REP g.leg1 \<Rightarrow> tab\<^sub>1 ?f \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f =
(?f \<star> \<theta>') \<cdot> \<a>[?f, tab\<^sub>0 ?f, w'] \<cdot> (f.tab \<star> w') \<cdot> \<nu>'"
using f.apex_unique_up_to_equivalence
[of "REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f"
"Maps.REP g.leg0" "Maps.REP g.leg1"]
by simp
obtain w w' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'
where 4: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src (tab\<^sub>0 ?f) \<rightarrow> src (Maps.REP g.leg0)\<guillemotright> \<and>
\<guillemotleft>w' : src (Maps.REP g.leg0) \<rightarrow> src (tab\<^sub>0 ?f)\<guillemotright> \<and>
\<guillemotleft>\<theta> : Maps.REP g.leg0 \<star> w \<Rightarrow> tab\<^sub>0 ?f\<guillemotright> \<and>
\<guillemotleft>\<nu> : tab\<^sub>1 ?f \<Rightarrow> Maps.REP g.leg1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
f.tab = (?f \<star> \<theta>) \<cdot> \<a>[?f, Maps.REP g.leg0, w] \<cdot>
(REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : tab\<^sub>0 ?f \<star> w' \<Rightarrow> Maps.REP g.leg0\<guillemotright> \<and>
\<guillemotleft>\<nu>' : Maps.REP g.leg1 \<Rightarrow> tab\<^sub>1 ?f \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f =
(?f \<star> \<theta>') \<cdot> \<a>[?f, tab\<^sub>0 ?f, w'] \<cdot> (f.tab \<star> w') \<cdot> \<nu>'"
using 3 by meson
hence w\<theta>\<nu>: "equivalence_map w \<and> \<guillemotleft>w : src (tab\<^sub>0 ?f) \<rightarrow> src (Maps.REP g.leg0)\<guillemotright> \<and>
\<guillemotleft>\<theta> : Maps.REP g.leg0 \<star> w \<Rightarrow> tab\<^sub>0 ?f\<guillemotright> \<and>
\<guillemotleft>\<nu> : tab\<^sub>1 ?f \<Rightarrow> Maps.REP g.leg1 \<star> w\<guillemotright> \<and> iso \<nu>"
using equivalence_map_def quasi_inverses_def quasi_inverses_symmetric
by meson
let ?W = "\<lparr>Chn = \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>, Dom = Dom (SPN ?f), Cod = Dom g\<rparr>"
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 \<open>Dom ?W\<close>
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 \<star> (Maps.REP g.leg0)\<^sup>*)"
assume 1: "\<guillemotleft>Maps.REP g.leg1 : Maps.Dom g.apex \<rightarrow> Maps.Cod g.leg1\<guillemotright>"
have "\<guillemotleft>?f : src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*))
\<rightarrow> Maps.Cod g.leg1\<guillemotright> \<and>
is_left_adjoint ?f \<and> \<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk> = \<lbrakk>?f\<rbrakk>"
using 1 by simp
thus "\<exists>f. \<guillemotleft>f : src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*))
\<rightarrow> Maps.Cod g.leg1\<guillemotright> \<and>
is_left_adjoint f \<and>
\<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk> = \<lbrakk>f\<rbrakk>"
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 \<star> (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 \<open>Cod ?W\<close>
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\<theta>\<nu> 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\<theta>\<nu> Maps.dom_char by simp
also have "... = Dom_W.apex"
proof -
have "src w = src (tab\<^sub>0 ?f)"
using w\<theta>\<nu> 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\<theta>\<nu> Maps.cod_char by simp
also have "... = Cod_W.apex"
proof -
have "trg w = src (Maps.REP g.leg0)"
using w\<theta>\<nu> 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 \<odot> Chn ?W = Dom_W.leg0"
proof -
have "Cod_W.leg0 \<odot> Chn ?W = g.leg0 \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
by simp
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using g.dom.leg_simps(1) Maps.CLS_REP [of g.leg0]
by simp
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg0 \<star> w\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (Maps.REP g.leg0)"
by fast
moreover have "is_left_adjoint w"
using w\<theta>\<nu> equivalence_is_adjoint by simp
moreover have "Maps.REP g.leg0 \<star> w \<cong> Maps.REP g.leg0 \<star> w"
using w\<theta>\<nu> 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\<theta>\<nu> Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint
by blast
qed
also have "... = \<lbrakk>\<lbrakk>tab\<^sub>0 ?f\<rbrakk>\<rbrakk>"
proof -
have "iso \<theta>"
proof -
have "is_left_adjoint (Maps.REP g.leg0 \<star> w)"
using w\<theta>\<nu> 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\<theta>\<nu> BS3 by blast
qed
thus ?thesis
using w\<theta>\<nu> 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 \<odot> Chn ?W = Dom_W.leg1"
proof -
have "Cod_W.leg1 \<odot> Chn ?W = g.leg1 \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
by simp
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using g.dom.leg_simps(3) Maps.CLS_REP by presburger
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg1 \<star> w\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (Maps.REP g.leg1)"
by fast
moreover have "is_left_adjoint w"
using w\<theta>\<nu> equivalence_is_adjoint by simp
moreover have "Maps.REP g.leg1 \<star> w \<cong> Maps.REP g.leg1 \<star> w"
using w\<theta>\<nu> 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\<theta>\<nu> Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint
by blast
qed
also have "... = \<lbrakk>\<lbrakk>tab\<^sub>1 ?f\<rbrakk>\<rbrakk>"
proof -
have "ide (Maps.REP g.leg1 \<star> w)"
using 2 w\<theta>\<nu> equivalence_map_is_ide by auto
moreover have "Maps.REP g.leg1 \<star> w \<cong>
tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)"
using w\<theta>\<nu> 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
\<open>\<lparr>Leg0 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))
(src (Maps.REP g.leg0)\<^sup>*)
(iso_class (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*))),
Leg1 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))
(Maps.Cod g.leg1)
(iso_class (tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))\<rparr>\<close>
using W Span.arr_char
by (simp add: arrow_of_spans_def)
interpret Cod_W: span_in_category Maps.comp \<open>Cod ?W\<close>
using W Span.arr_char
by (simp add: arrow_of_spans_def)
show "Span.dom ?W = SPN ?f"
proof -
have "Span.dom ?W =
\<lparr>Chn = Dom_W.apex,
Dom = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>,
Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>\<rparr>,
Cod = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>,
Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>\<rparr>\<rparr>"
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 \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
proof -
have "Maps.arr \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk> \<and> w \<in> Maps.Map \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk> \<and> equivalence_map w"
proof (intro conjI)
show 1: "Maps.arr \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using w\<theta>\<nu> Maps.CLS_in_hom equivalence_is_adjoint by blast
show "equivalence_map w"
using w\<theta>\<nu> by blast
show "w \<in> Maps.Map \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using 1 w\<theta>\<nu> 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\<theta>\<nu> 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 "\<And>r s \<tau>. \<lbrakk>ide r; ide s; src r = src s; trg r = trg s; Span.in_hom \<tau> (SPN r) (SPN s)\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright> \<and> SPN \<mu> = \<tau>"
proof -
fix r s \<tau>
assume r: "ide r" and s: "ide s"
assume src_eq: "src r = src s" and trg_eq: "trg r = trg s"
assume \<tau>: "Span.in_hom \<tau> (SPN r) (SPN s)"
interpret \<tau>: arrow_of_spans Maps.comp \<tau>
using \<tau> Span.arr_char by auto
interpret r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r
using r by unfold_locales auto
interpret s: identity_in_bicategory_of_spans V H \<a> \<i> src trg s
using s by unfold_locales auto
interpret s: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close> s
using s.is_arrow_of_tabulations_in_maps by simp
have \<tau>_dom_leg0_eq: "\<tau>.dom.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk>"
using \<tau> Span.dom_char SPN_def [of r] by auto
have \<tau>_dom_leg1_eq: "\<tau>.dom.leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>"
using \<tau> Span.dom_char SPN_def [of r] by auto
have \<tau>_cod_leg0_eq: "\<tau>.cod.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 s\<rbrakk>\<rbrakk>"
using \<tau> Span.cod_char SPN_def [of s] by auto
have \<tau>_cod_leg1_eq: "\<tau>.cod.leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 s\<rbrakk>\<rbrakk>"
using \<tau> Span.cod_char SPN_def [of s] by auto
have 1: "tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<cong> tab\<^sub>0 r"
proof -
have "tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<cong> Maps.REP \<tau>.cod.leg0 \<star> Maps.REP \<tau>.chine"
proof -
have "Maps.REP \<tau>.cod.leg0 \<cong> tab\<^sub>0 s"
using \<tau>_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 \<tau>.chine)"
using Maps.seq_char [of "\<lbrakk>\<lbrakk>tab\<^sub>0 s\<rbrakk>\<rbrakk>" \<tau>.chine] \<tau>.cod.leg_simps(1)
\<tau>.leg0_commutes \<tau>_cod_leg0_eq
by auto
ultimately show ?thesis
using hcomp_isomorphic_ide [of "tab\<^sub>0 s" "Maps.REP \<tau>.cod.leg0" "Maps.REP \<tau>.chine"]
isomorphic_symmetric
by fastforce
qed
also have "... \<cong> Maps.REP \<tau>.dom.leg0"
proof -
have "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg0\<rbrakk>\<rbrakk>"
using \<tau>.leg0_commutes Maps.CLS_REP \<tau>.chine_simps(1)
\<tau>.cod.leg_simps(1) \<tau>.dom.leg_simps(1)
by presburger
hence "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg0 \<star> Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg0\<rbrakk>\<rbrakk>"
using Maps.comp_CLS [of "Maps.REP \<tau>.cod.leg0" "Maps.REP \<tau>.chine"
"Maps.REP \<tau>.cod.leg0 \<star> Maps.REP \<tau>.chine"]
isomorphic_reflexive
by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2) Maps.REP_simps(2-3)
\<tau>.chine_in_hom \<tau>.cod.leg_in_hom(1) \<tau>.dom.leg_simps(1) \<tau>.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 "... \<cong> tab\<^sub>0 r"
using \<tau>_dom_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.satisfies_T0
by presburger
finally show ?thesis by blast
qed
obtain \<theta> where \<theta>: "\<guillemotleft>\<theta> : tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<Rightarrow> tab\<^sub>0 r\<guillemotright> \<and> iso \<theta>"
using 1 by blast
have 2: "tab\<^sub>1 s \<star> Maps.REP \<tau>.chine \<cong> tab\<^sub>1 r"
proof -
have "tab\<^sub>1 s \<star> Maps.REP \<tau>.chine \<cong> Maps.REP \<tau>.cod.leg1 \<star> Maps.REP \<tau>.chine"
proof -
have "Maps.REP \<tau>.cod.leg1 \<cong> tab\<^sub>1 s"
using \<tau>_cod_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.leg1_is_map
by presburger
moreover have "src (Maps.REP \<tau>.cod.leg1) = trg (Maps.REP \<tau>.chine)"
using Maps.seq_char by auto
ultimately show ?thesis
using hcomp_isomorphic_ide [of "Maps.REP \<tau>.cod.leg1" "tab\<^sub>1 s" "Maps.REP \<tau>.chine"]
isomorphic_symmetric
by fastforce
qed
also have "... \<cong> Maps.REP \<tau>.dom.leg1"
proof -
have "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg1\<rbrakk>\<rbrakk>"
using \<tau>.leg1_commutes Maps.CLS_REP \<tau>.chine_simps(1)
\<tau>.cod.leg_simps(3) \<tau>.dom.leg_simps(3)
by presburger
hence "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg1 \<star> Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg1\<rbrakk>\<rbrakk>"
using Maps.comp_CLS [of "Maps.REP \<tau>.cod.leg1" "Maps.REP \<tau>.chine"
"Maps.REP \<tau>.cod.leg1 \<star> Maps.REP \<tau>.chine"]
isomorphic_reflexive
by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2)
Maps.REP_simps(2) Maps.REP_simps(3) \<tau>.chine_in_hom \<tau>.cod.leg_in_hom(2)
\<tau>.dom.leg_simps(3) \<tau>.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 "... \<cong> tab\<^sub>1 r"
using \<tau>_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 \<nu> where \<nu>: "\<guillemotleft>\<nu> : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 s \<star> Maps.REP \<tau>.chine\<guillemotright> \<and> iso \<nu>"
using 2 isomorphic_symmetric by blast
define \<Delta>
where "\<Delta> \<equiv> (s \<star> \<theta>) \<cdot> \<a>[s, tab\<^sub>0 s, Maps.REP \<tau>.chine] \<cdot> (s.tab \<star> Maps.REP \<tau>.chine) \<cdot> \<nu>"
have \<Delta>: "\<guillemotleft>\<Delta> : tab\<^sub>1 r \<Rightarrow> s \<star> tab\<^sub>0 r\<guillemotright>"
proof (unfold \<Delta>_def, intro comp_in_homI)
show "\<guillemotleft>\<nu> : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 s \<star> Maps.REP \<tau>.chine\<guillemotright>"
using \<nu> by simp
show 3: "\<guillemotleft>s.tab \<star> Maps.REP \<tau>.chine :
tab\<^sub>1 s \<star> Maps.REP \<tau>.chine \<Rightarrow> (s \<star> tab\<^sub>0 s) \<star> Maps.REP \<tau>.chine\<guillemotright>"
apply (intro hcomp_in_vhom)
apply auto
using "1" by fastforce
show "\<guillemotleft>\<a>[s, tab\<^sub>0 s, Maps.REP \<tau>.chine] :
(s \<star> tab\<^sub>0 s) \<star> Maps.REP \<tau>.chine \<Rightarrow> s \<star> tab\<^sub>0 s \<star> Maps.REP \<tau>.chine\<guillemotright>"
using s assoc_in_hom [of s "tab\<^sub>0 s" "Maps.REP \<tau>.chine"]
by (metis (no_types, lifting) Maps.ide_REP 3 \<tau>.chine_simps(1) hcomp_in_vhomE
ideD(2) ideD(3) s.ide_u s.tab_simps(2) s.u_simps(3))
show "\<guillemotleft>s \<star> \<theta> : s \<star> tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<Rightarrow> s \<star> tab\<^sub>0 r\<guillemotright>"
using 1 s \<theta> isomorphic_implies_hpar(4) src_eq by auto
qed
define \<mu> where "\<mu> \<equiv> r.T0.trnr\<^sub>\<epsilon> s \<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab)"
have \<mu>: "\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>"
proof (unfold \<mu>_def, intro comp_in_homI)
show "\<guillemotleft>inv (r.T0.trnr\<^sub>\<epsilon> r r.tab) : r \<Rightarrow> tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*\<guillemotright>"
using r.yields_isomorphic_representation by fastforce
show "\<guillemotleft>r.T0.trnr\<^sub>\<epsilon> s \<Delta> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> s\<guillemotright>"
using s \<Delta> src_eq r.T0.adjoint_transpose_right(2) [of s "tab\<^sub>1 r"] by auto
qed
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg r s \<mu>
using \<mu> by unfold_locales auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close> \<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
have \<Delta>_eq: "\<Delta> = \<mu>.\<Delta>"
proof -
have "r.T0.trnr\<^sub>\<epsilon> s \<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab) =
r.T0.trnr\<^sub>\<epsilon> s \<mu>.\<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab)"
using \<mu> \<mu>_def \<mu>.\<mu>_in_terms_of_\<Delta> by auto
hence "r.T0.trnr\<^sub>\<epsilon> s \<Delta> = r.T0.trnr\<^sub>\<epsilon> s \<mu>.\<Delta>"
using r s \<Delta> r.T0.adjoint_transpose_right(2) r.yields_isomorphic_representation
iso_inv_iso iso_is_retraction retraction_is_epi epiE
by (metis \<mu>.in_hom \<mu>_def arrI)
thus ?thesis
using \<Delta> \<mu>.\<Delta>_in_hom(2) src_eq r.T0.adjoint_transpose_right(6)
bij_betw_imp_inj_on
[of "r.T0.trnr\<^sub>\<epsilon> s" "hom (tab\<^sub>1 r) (s \<star> tab\<^sub>0 r)" "hom (tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*) s"]
inj_on_def [of "r.T0.trnr\<^sub>\<epsilon> s" "hom (tab\<^sub>1 r) (s \<star> tab\<^sub>0 r)"]
by simp
qed
have "\<mu>.is_induced_map (Maps.REP \<tau>.chine)"
using \<theta> \<nu> \<Delta>_eq \<Delta>_def \<mu>.is_induced_map_iff \<tau>.chine_simps(1) Maps.ide_REP by blast
hence 3: "Maps.REP \<tau>.chine \<cong> \<mu>.chine"
using \<mu>.chine_is_induced_map \<mu>.induced_map_unique by simp
have "SPN \<mu> = \<tau>"
proof (intro Span.arr_eqI)
show "Span.par (SPN \<mu>) \<tau>"
using \<mu> \<tau> SPN_in_hom
by (metis (no_types, lifting) SPN.preserves_cod SPN.preserves_dom Span.in_homE
in_homE)
show "Chn (SPN \<mu>) = \<tau>.chine"
proof -
have "Chn (SPN \<mu>) = \<lbrakk>\<lbrakk>spn \<mu>\<rbrakk>\<rbrakk>"
using \<mu> SPN_def spn_def by auto
also have "... = \<lbrakk>\<lbrakk>\<mu>.chine\<rbrakk>\<rbrakk>"
using \<mu> spn_def by fastforce
also have "... = \<lbrakk>\<lbrakk>Maps.REP \<tau>.chine\<rbrakk>\<rbrakk>"
using 3 isomorphic_symmetric Maps.CLS_eqI iso_class_eqI isomorphic_implies_hpar(3)
isomorphic_implies_hpar(4)
by auto
also have "... = \<tau>.chine"
using Maps.CLS_REP \<tau>.chine_simps(1) by blast
finally show ?thesis by blast
qed
qed
thus "\<exists>\<mu>. \<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright> \<and> SPN \<mu> = \<tau>"
using \<mu> by auto
qed
qed
theorem SPN_is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>"
..
text \<open>
We have completed the proof of the second half of the main result (CKS Theorem 4):
\<open>B\<close> is biequivalent (via \<open>SPN\<close>) to \<open>Span(Maps(B))\<close>.
\<close>
corollary
shows "equivalent_bicategories V H \<a> \<i> 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,3959 +1,3959 @@
(* Title: Coherence
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Coherence"
theory Coherence
imports Bicategory
begin
text \<open>
\sloppypar
In this section, we generalize to bicategories the proof of the Coherence Theorem
that we previously gave for monoidal categories
(see \<open>MonoidalCategory.evaluation_map.coherence\<close> 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 \<open>t\<close> to a
``canonical'' term \<open>t\<^bold>\<down>\<close> that connects \<open>t\<close> 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''.
\<close>
subsection "Bicategorical Language"
text \<open>
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 \<open>C\<close> and the syntax contained a constructor
to form a primitive term corresponding to each arrow of \<open>C\<close>.
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''.
\<close>
locale bicategorical_language =
category V +
horizontal_homs V src trg
for V :: "'a comp" (infixr "\<cdot>" 55)
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
begin
text \<open>
Constructor \<open>Prim\<^sub>0\<close> is used to construct ``formal objects'' and \<open>Prim\<close> is used to
construct primitive terms that are not formal objects.
\<close>
datatype (discs_sels) 't "term" =
Prim\<^sub>0 't ("\<^bold>\<langle>_\<^bold>\<rangle>\<^sub>0")
| Prim 't ("\<^bold>\<langle>_\<^bold>\<rangle>")
| Hcomp "'t term" "'t term" (infixr "\<^bold>\<star>" 53)
| Vcomp "'t term" "'t term" (infixr "\<^bold>\<cdot>" 55)
| Lunit "'t term" ("\<^bold>\<l>\<^bold>[_\<^bold>]")
| Lunit' "'t term" ("\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[_\<^bold>]")
| Runit "'t term" ("\<^bold>\<r>\<^bold>[_\<^bold>]")
| Runit' "'t term" ("\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[_\<^bold>]")
| Assoc "'t term" "'t term" "'t term" ("\<^bold>\<a>\<^bold>[_, _, _\<^bold>]")
| Assoc' "'t term" "'t term" "'t term" ("\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[_, _, _\<^bold>]")
text \<open>
We define formal domain, codomain, source, and target functions on terms.
\<close>
primrec Src :: "'a term \<Rightarrow> 'a term"
where "Src \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
| "Src \<^bold>\<langle>\<mu>\<^bold>\<rangle> = \<^bold>\<langle>src \<mu>\<^bold>\<rangle>\<^sub>0"
| "Src (t \<^bold>\<star> u) = Src u"
| "Src (t \<^bold>\<cdot> u) = Src t"
| "Src \<^bold>\<l>\<^bold>[t\<^bold>] = Src t"
| "Src \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Src t"
| "Src \<^bold>\<r>\<^bold>[t\<^bold>] = Src t"
| "Src \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Src t"
| "Src \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = Src v"
| "Src \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = Src v"
primrec Trg :: "'a term \<Rightarrow> 'a term"
where "Trg \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
| "Trg \<^bold>\<langle>\<mu>\<^bold>\<rangle> = \<^bold>\<langle>trg \<mu>\<^bold>\<rangle>\<^sub>0"
| "Trg (t \<^bold>\<star> u) = Trg t"
| "Trg (t \<^bold>\<cdot> u) = Trg t"
| "Trg \<^bold>\<l>\<^bold>[t\<^bold>] = Trg t"
| "Trg \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Trg t"
| "Trg \<^bold>\<r>\<^bold>[t\<^bold>] = Trg t"
| "Trg \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Trg t"
| "Trg \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = Trg t"
| "Trg \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = Trg t"
primrec Dom :: "'a term \<Rightarrow> 'a term"
where "Dom \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
| "Dom \<^bold>\<langle>\<mu>\<^bold>\<rangle> = \<^bold>\<langle>dom \<mu>\<^bold>\<rangle>"
| "Dom (t \<^bold>\<star> u) = Dom t \<^bold>\<star> Dom u"
| "Dom (t \<^bold>\<cdot> u) = Dom u"
| "Dom \<^bold>\<l>\<^bold>[t\<^bold>] = Trg t \<^bold>\<star> Dom t"
| "Dom \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Dom t"
| "Dom \<^bold>\<r>\<^bold>[t\<^bold>] = Dom t \<^bold>\<star> Src t"
| "Dom \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Dom t"
| "Dom \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = (Dom t \<^bold>\<star> Dom u) \<^bold>\<star> Dom v"
| "Dom \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = Dom t \<^bold>\<star> (Dom u \<^bold>\<star> Dom v)"
primrec Cod :: "'a term \<Rightarrow> 'a term"
where "Cod \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
| "Cod \<^bold>\<langle>\<mu>\<^bold>\<rangle> = \<^bold>\<langle>cod \<mu>\<^bold>\<rangle>"
| "Cod (t \<^bold>\<star> u) = Cod t \<^bold>\<star> Cod u"
| "Cod (t \<^bold>\<cdot> u) = Cod t"
| "Cod \<^bold>\<l>\<^bold>[t\<^bold>] = Cod t"
| "Cod \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Trg t \<^bold>\<star> Cod t"
| "Cod \<^bold>\<r>\<^bold>[t\<^bold>] = Cod t"
| "Cod \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Cod t \<^bold>\<star> Src t"
| "Cod \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = Cod t \<^bold>\<star> (Cod u \<^bold>\<star> Cod v)"
| "Cod \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Cod t \<^bold>\<star> Cod u) \<^bold>\<star> Cod v"
text \<open>
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.
\<close>
primrec Arr :: "'a term \<Rightarrow> bool"
where "Arr \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = obj \<mu>"
| "Arr \<^bold>\<langle>\<mu>\<^bold>\<rangle> = arr \<mu>"
| "Arr (t \<^bold>\<star> u) = (Arr t \<and> Arr u \<and> Src t = Trg u)"
| "Arr (t \<^bold>\<cdot> u) = (Arr t \<and> Arr u \<and> Dom t = Cod u)"
| "Arr \<^bold>\<l>\<^bold>[t\<^bold>] = Arr t"
| "Arr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Arr t"
| "Arr \<^bold>\<r>\<^bold>[t\<^bold>] = Arr t"
| "Arr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Arr t"
| "Arr \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = (Arr t \<and> Arr u \<and> Arr v \<and> Src t = Trg u \<and> Src u = Trg v)"
| "Arr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Arr t \<and> Arr u \<and> Arr v \<and> Src t = Trg u \<and> Src u = Trg v)"
primrec Ide :: "'a term \<Rightarrow> bool"
where "Ide \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = obj \<mu>"
| "Ide \<^bold>\<langle>\<mu>\<^bold>\<rangle> = ide \<mu>"
| "Ide (t \<^bold>\<star> u) = (Ide t \<and> Ide u \<and> Src t = Trg u)"
| "Ide (t \<^bold>\<cdot> u) = False"
| "Ide \<^bold>\<l>\<^bold>[t\<^bold>] = False"
| "Ide \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False"
| "Ide \<^bold>\<r>\<^bold>[t\<^bold>] = False"
| "Ide \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False"
| "Ide \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = False"
| "Ide \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = False"
primrec Obj :: "'a term \<Rightarrow> bool"
where "Obj \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = obj \<mu>"
| "Obj \<^bold>\<langle>\<mu>\<^bold>\<rangle> = False"
| "Obj (t \<^bold>\<star> u) = False"
| "Obj (t \<^bold>\<cdot> u) = False"
| "Obj \<^bold>\<l>\<^bold>[t\<^bold>] = False"
| "Obj \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False"
| "Obj \<^bold>\<r>\<^bold>[t\<^bold>] = False"
| "Obj \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False"
| "Obj \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = False"
| "Obj \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = False"
abbreviation HSeq :: "'a term \<Rightarrow> 'a term \<Rightarrow> bool"
where "HSeq t u \<equiv> Arr t \<and> Arr u \<and> Src t = Trg u"
abbreviation VSeq :: "'a term \<Rightarrow> 'a term \<Rightarrow> bool"
where "VSeq t u \<equiv> Arr t \<and> Arr u \<and> Dom t = Cod u"
abbreviation HPar :: "'a term => 'a term \<Rightarrow> bool"
where "HPar t u \<equiv> Arr t \<and> Arr u \<and> Src t = Src u \<and> Trg t = Trg u"
abbreviation VPar :: "'a term => 'a term \<Rightarrow> bool"
where "VPar t u \<equiv> Arr t \<and> Arr u \<and> Dom t = Dom u \<and> Cod t = Cod u"
abbreviation HHom :: "'a term \<Rightarrow> 'a term \<Rightarrow> 'a term set"
where "HHom a b \<equiv> { t. Arr t \<and> Src t = a \<and> Trg t = b }"
abbreviation VHom :: "'a term \<Rightarrow> 'a term \<Rightarrow> 'a term set"
where "VHom f g \<equiv> { t. Arr t \<and> Dom t = f \<and> 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 \<Longrightarrow> Src (Src t) = Src t"
by (induct t) auto
lemma Trg_Trg [simp]:
shows "Arr t \<Longrightarrow> Trg (Trg t) = Trg t"
by (induct t) auto
lemma Src_Trg [simp]:
shows "Arr t \<Longrightarrow> Src (Trg t) = Trg t"
by (induct t) auto
lemma Trg_Src [simp]:
shows "Arr t \<Longrightarrow> Trg (Src t) = Src t"
by (induct t) auto
lemma Dom_Src [simp]:
shows "Arr t \<Longrightarrow> Dom (Src t) = Src t"
by (induct t) auto
lemma Dom_Trg [simp]:
shows "Arr t \<Longrightarrow> Dom (Trg t) = Trg t"
by (induct t) auto
lemma Cod_Src [simp]:
shows "Arr t \<Longrightarrow> Cod (Src t) = Src t"
by (induct t) auto
lemma Cod_Trg [simp]:
shows "Arr t \<Longrightarrow> Cod (Trg t) = Trg t"
by (induct t) auto
lemma Src_Dom_Cod:
shows "Arr t \<Longrightarrow> Src (Dom t) = Src t \<and> Src (Cod t) = Src t"
using src_dom src_cod by (induct t) auto
lemma Src_Dom [simp]:
shows "Arr t \<Longrightarrow> Src (Dom t) = Src t"
using Src_Dom_Cod by blast
lemma Src_Cod [simp]:
shows "Arr t \<Longrightarrow> Src (Cod t) = Src t"
using Src_Dom_Cod by blast
lemma Trg_Dom_Cod:
shows "Arr t \<Longrightarrow> Trg (Dom t) = Trg t \<and> Trg (Cod t) = Trg t"
using trg_dom trg_cod by (induct t) auto
lemma Trg_Dom [simp]:
shows "Arr t \<Longrightarrow> Trg (Dom t) = Trg t"
using Trg_Dom_Cod by blast
lemma Trg_Cod [simp]:
shows "Arr t \<Longrightarrow> Trg (Cod t) = Trg t"
using Trg_Dom_Cod by blast
lemma VSeq_implies_HPar:
shows "VSeq t u \<Longrightarrow> 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 \<Longrightarrow> Dom (Dom t) = Dom t"
by (induct t, auto)
lemma Cod_Cod [simp]:
shows "Arr t \<Longrightarrow> Cod (Cod t) = Cod t"
by (induct t, auto)
lemma Dom_Cod [simp]:
shows "Arr t \<Longrightarrow> Dom (Cod t) = Cod t"
by (induct t, auto)
lemma Cod_Dom [simp]:
shows "Arr t \<Longrightarrow> Cod (Dom t) = Dom t"
by (induct t, auto)
lemma Obj_implies_Ide (*[simp]*):
shows "Obj t \<Longrightarrow> Ide t"
by (induct t) auto
lemma Ide_implies_Arr [simp]:
shows "Ide t \<Longrightarrow> Arr t"
by (induct t, auto)
lemma Dom_Ide:
shows "Ide t \<Longrightarrow> Dom t = t"
by (induct t, auto)
lemma Cod_Ide:
shows "Ide t \<Longrightarrow> Cod t = t"
by (induct t, auto)
lemma Obj_Src [simp]:
shows "Arr t \<Longrightarrow> Obj (Src t)"
by (induct t) auto
lemma Obj_Trg [simp]:
shows "Arr t \<Longrightarrow> Obj (Trg t)"
by (induct t) auto
lemma Ide_Dom [simp]:
shows "Arr t \<Longrightarrow> Ide (Dom t)"
using Obj_implies_Ide
by (induct t, auto)
lemma Ide_Cod [simp]:
shows "Arr t \<Longrightarrow> Ide (Cod t)"
using Obj_implies_Ide
by (induct t, auto)
lemma Arr_in_Hom:
assumes "Arr t"
shows "t \<in> HHom (Src t) (Trg t)" and "t \<in> VHom (Dom t) (Cod t)"
proof -
have 1: "Arr t \<Longrightarrow> t \<in> HHom (Src t) (Trg t) \<and> t \<in> VHom (Dom t) (Cod t)"
by (induct t, auto)
show "t \<in> HHom (Src t) (Trg t)" using assms 1 by simp
show "t \<in> VHom (Dom t) (Cod t)" using assms 1 by simp
qed
lemma Ide_in_Hom:
assumes "Ide t"
shows "t \<in> HHom (Src t) (Trg t)" and "t \<in> VHom t t"
proof -
have 1: "Ide t \<Longrightarrow> t \<in> HHom (Src t) (Trg t) \<and> t \<in> VHom t t"
by (induct t, auto)
show "t \<in> HHom (Src t) (Trg t)" using assms 1 by simp
show "t \<in> VHom t t" using assms 1 by simp
qed
lemma Obj_in_Hom:
assumes "Obj t"
shows "t \<in> HHom t t" and "t \<in> VHom t t"
proof -
have 1: "Obj t \<Longrightarrow> t \<in> HHom t t \<and> t \<in> VHom t t"
by (induct t, auto)
show "t \<in> HHom t t" using assms 1 by simp
show "t \<in> VHom t t" using assms 1 by simp
qed
text \<open>
A formal arrow is ``canonical'' if the only primitive arrows used in its construction
are objects and identities.
\<close>
primrec Can :: "'a term \<Rightarrow> bool"
where "Can \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = obj \<mu>"
| "Can \<^bold>\<langle>\<mu>\<^bold>\<rangle> = ide \<mu>"
| "Can (t \<^bold>\<star> u) = (Can t \<and> Can u \<and> Src t = Trg u)"
| "Can (t \<^bold>\<cdot> u) = (Can t \<and> Can u \<and> Dom t = Cod u)"
| "Can \<^bold>\<l>\<^bold>[t\<^bold>] = Can t"
| "Can \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Can t"
| "Can \<^bold>\<r>\<^bold>[t\<^bold>] = Can t"
| "Can \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Can t"
| "Can \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = (Can t \<and> Can u \<and> Can v \<and> Src t = Trg u \<and> Src u = Trg v)"
| "Can \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Can t \<and> Can u \<and> Can v \<and> Src t = Trg u \<and> Src u = Trg v)"
lemma Ide_implies_Can:
shows "Ide t \<Longrightarrow> Can t"
by (induct t, auto)
lemma Can_implies_Arr:
shows "Can t \<Longrightarrow> Arr t"
by (induct t, auto)
text \<open>
Canonical arrows can be formally inverted.
\<close>
primrec Inv :: "'a term \<Rightarrow> 'a term"
where "Inv \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
| "Inv \<^bold>\<langle>\<mu>\<^bold>\<rangle> = \<^bold>\<langle>inv \<mu>\<^bold>\<rangle>"
| "Inv (t \<^bold>\<star> u) = (Inv t \<^bold>\<star> Inv u)"
| "Inv (t \<^bold>\<cdot> u) = (Inv u \<^bold>\<cdot> Inv t)"
| "Inv \<^bold>\<l>\<^bold>[t\<^bold>] = \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]"
| "Inv \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<^bold>\<l>\<^bold>[Inv t\<^bold>]"
| "Inv \<^bold>\<r>\<^bold>[t\<^bold>] = \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]"
| "Inv \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<^bold>\<r>\<^bold>[Inv t\<^bold>]"
| "Inv \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Inv t, Inv u, Inv v\<^bold>]"
| "Inv \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = \<^bold>\<a>\<^bold>[Inv t, Inv u, Inv v\<^bold>]"
lemma Src_Inv [simp]:
shows "Can t \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> Dom (Inv t) = Cod t"
by (induct t, auto)
lemma Cod_Inv [simp]:
shows "Can t \<Longrightarrow> Cod (Inv t) = Dom t"
by (induct t, auto)
lemma Inv_preserves_Ide:
shows "Ide t \<Longrightarrow> Ide (Inv t)"
using Ide_implies_Can by (induct t, auto)
lemma Can_Inv [simp]:
shows "Can t \<Longrightarrow> Can (Inv t)"
by (induct t, auto)
lemma Inv_in_Hom [intro]:
assumes "Can t"
shows "Inv t \<in> HHom (Src t) (Trg t)" and "Inv t \<in> 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 \<open>
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 \<open>\<^bold>\<star>\<close> is used as
the list constructor.
\<close>
fun Nml :: "'a term \<Rightarrow> bool"
where "Nml \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = obj \<mu>"
| "Nml \<^bold>\<langle>\<mu>\<^bold>\<rangle> = arr \<mu>"
| "Nml (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<star> u) = (arr \<nu> \<and> Nml u \<and> \<not> is_Prim\<^sub>0 u \<and> \<^bold>\<langle>src \<nu>\<^bold>\<rangle>\<^sub>0 = Trg u)"
| "Nml _ = False"
lemma Nml_HcompD:
assumes "Nml (t \<^bold>\<star> u)"
shows "\<^bold>\<langle>un_Prim t\<^bold>\<rangle> = t" and "arr (un_Prim t)" and "Nml t" and "Nml u"
and "\<not> is_Prim\<^sub>0 u" and "\<^bold>\<langle>src (un_Prim t)\<^bold>\<rangle>\<^sub>0 = Trg u" and "Src t = Trg u"
proof -
have 1: "t = \<^bold>\<langle>un_Prim t\<^bold>\<rangle> \<and> arr (un_Prim t) \<and> Nml t \<and> Nml u \<and> \<not> is_Prim\<^sub>0 u \<and>
\<^bold>\<langle>src (un_Prim t)\<^bold>\<rangle>\<^sub>0 = Trg u"
using assms by (cases t; simp; cases u; simp)
show "\<^bold>\<langle>un_Prim t\<^bold>\<rangle> = 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 "\<not> is_Prim\<^sub>0 u" using 1 by simp
show "\<^bold>\<langle>src (un_Prim t)\<^bold>\<rangle>\<^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 \<Longrightarrow> Arr t"
by (induct t, auto simp add: Nml_HcompD)
lemma Nml_Src [simp]:
shows "Nml t \<Longrightarrow> Nml (Src t)"
apply (induct t, simp_all)
using Nml_HcompD by metis
lemma Nml_Trg [simp]:
shows "Nml t \<Longrightarrow> Nml (Trg t)"
apply (induct t, simp_all)
using Nml_HcompD by metis
lemma Nml_Dom [simp]:
shows "Nml t \<Longrightarrow> 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>\<star> v)"
show "Nml (Dom u \<^bold>\<star> Dom v)"
proof -
have 1: "is_Prim (Dom u) \<and> arr (un_Prim (Dom u)) \<and> Dom u = \<^bold>\<langle>dom (un_Prim u)\<^bold>\<rangle>"
using uv by (cases u; simp; cases v, simp_all)
have 2: "Nml v \<and> \<not> is_Prim\<^sub>0 v \<and> \<not> is_Vcomp v \<and> \<not> is_Lunit' v \<and> \<not> 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) \<and> \<not> is_Prim\<^sub>0 v"
using 2 I2 by (cases v, simp_all)
moreover have "\<^bold>\<langle>src (dom (un_Prim u))\<^bold>\<rangle>\<^sub>0 = Trg (Dom v)"
proof -
have "Trg (Dom v) = Src (Dom u)"
using uv Nml_implies_Arr by fastforce
also have "... = \<^bold>\<langle>src (dom (un_Prim u))\<^bold>\<rangle>\<^sub>0"
using 1 by fastforce
finally show ?thesis by argo
qed
moreover have "\<not> 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 \<Longrightarrow> 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>\<star> v)"
show "Nml (Cod u \<^bold>\<star> Cod v)"
proof -
have 1: "is_Prim (Cod u) \<and> arr (un_Prim (Cod u)) \<and> Cod u = \<^bold>\<langle>cod (un_Prim u)\<^bold>\<rangle>"
using uv by (cases u; simp; cases v, simp_all)
have 2: "Nml v \<and> \<not> is_Prim\<^sub>0 v \<and> \<not> is_Vcomp v \<and> \<not> is_Lunit' v \<and> \<not> 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) \<and> \<not> is_Prim\<^sub>0 v"
using 2 I2 by (cases v, simp_all)
moreover have "\<^bold>\<langle>src (cod (un_Prim u))\<^bold>\<rangle>\<^sub>0 = Trg (Cod v)"
proof -
have "Trg (Cod v) = Src (Cod u)"
using uv Nml_implies_Arr by fastforce
also have "... = \<^bold>\<langle>src (cod (un_Prim u))\<^bold>\<rangle>\<^sub>0"
using 1 by fastforce
finally show ?thesis by argo
qed
moreover have "\<not> 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 \<and> Nml t \<Longrightarrow> Nml (Inv t)"
apply (induct t, simp_all)
proof -
fix u v
assume I1: "Nml u \<Longrightarrow> Nml (Inv u)"
assume I2: "Nml v \<Longrightarrow> Nml (Inv v)"
assume uv: "Can u \<and> Can v \<and> Src u = Trg v \<and> Nml (u \<^bold>\<star> v)"
show "Nml (Inv u \<^bold>\<star> Inv v)"
proof -
have u: "Arr u \<and> Can u" using uv Can_implies_Arr by blast
have v: "Arr v \<and> Can v" using uv Can_implies_Arr by blast
have 1: "\<^bold>\<langle>un_Prim u\<^bold>\<rangle> = u \<and> ide (un_Prim u) \<and> Nml u \<and> Nml v \<and> \<not> is_Prim\<^sub>0 v \<and>
\<^bold>\<langle>src (un_Prim u)\<^bold>\<rangle>\<^sub>0 = Trg v"
using uv Nml_HcompD [of u v] apply simp
using uv by (cases u, simp_all)
have 2: "\<^bold>\<langle>un_Prim (Inv u)\<^bold>\<rangle> = Inv u \<and> arr (un_Prim (Inv u)) \<and> Nml (Inv u)"
using 1 by (cases u; simp)
moreover have "Nml (Inv v) \<and> \<not> is_Prim\<^sub>0 (Inv v)"
using 1 I2 by (cases v, simp_all)
moreover have "\<^bold>\<langle>src (un_Prim (Inv u))\<^bold>\<rangle>\<^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 \<open>
The following function defines a horizontal composition for normal terms.
If such terms are regarded as lists, this is just (typed) list concatenation.
\<close>
fun HcompNml (infixr "\<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor>" 53)
where "\<^bold>\<langle>\<nu>\<^bold>\<rangle>\<^sub>0 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = u"
| "t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = t"
| "\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = \<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<star> u"
| "(t \<^bold>\<star> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
| "t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = undefined"
lemma HcompNml_Prim [simp]:
assumes "\<not> is_Prim\<^sub>0 t"
shows "\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> t = \<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<star> t"
using assms by (cases t, simp_all)
lemma HcompNml_term_Prim\<^sub>0 [simp]:
assumes "Src t = Trg \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
shows "t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = t"
using assms by (cases t, simp_all)
lemma HcompNml_Nml:
assumes "Nml (t \<^bold>\<star> u)"
shows "t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = t \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
and "Dom (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Dom t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u"
and "Cod (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Cod t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u"
and "Src (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u"
and "Trg (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Trg t"
proof -
have 0: "\<And>u. \<lbrakk> Nml t; Nml u; Src t = Trg u \<rbrakk> \<Longrightarrow>
Nml (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and> Dom (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Dom t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u \<and>
Cod (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Cod t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u \<and>
Src (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u \<and> Trg (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Trg t"
proof (induct t, simp_all add: Nml_HcompD(1-4))
fix \<nu> :: 'a and u :: "'a term"
assume \<nu>: "arr \<nu>"
assume u: "Nml u"
assume 1: "\<^bold>\<langle>src \<nu>\<^bold>\<rangle>\<^sub>0 = Trg u"
show "Nml (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and> Dom (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = \<^bold>\<langle>dom \<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u \<and>
Cod (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = \<^bold>\<langle>cod \<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u \<and>
Src (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u \<and> Trg (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = \<^bold>\<langle>trg \<nu>\<^bold>\<rangle>\<^sub>0"
using u \<nu> 1 by (cases u, simp_all)
next
fix u v w
assume I1: "\<And>x. Nml x \<Longrightarrow> Src v = Trg x \<Longrightarrow>
Nml (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) \<and> Dom (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom x \<and>
Cod (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Cod v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod x \<and>
Src (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Src x \<and> Trg (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Trg v"
assume I2: "\<And>x. Nml x \<Longrightarrow> Trg u = Trg x \<Longrightarrow>
Nml (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) \<and> Dom (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom x \<and>
Cod (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Cod w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod x \<and>
Src (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Src x \<and> Trg (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Trg w"
assume vw: "Nml (v \<^bold>\<star> w)"
assume u: "Nml u"
assume wu: "Src w = Trg u"
show "Nml ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and>
Dom ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = (Dom v \<^bold>\<star> Dom w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u \<and>
Cod ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = (Cod v \<^bold>\<star> Cod w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u \<and>
Src ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u \<and> Trg ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Trg v"
proof -
have v: "v = \<^bold>\<langle>un_Prim v\<^bold>\<rangle> \<and> Nml v"
using vw Nml_implies_Arr Nml_HcompD by metis
have w: "Nml w \<and> \<not> is_Prim\<^sub>0 w \<and> \<^bold>\<langle>src (un_Prim v)\<^bold>\<rangle>\<^sub>0 = Trg w"
using vw by (simp add: Nml_HcompD)
have "is_Prim\<^sub>0 u \<Longrightarrow> ?thesis" by (cases u; simp add: vw wu)
moreover have "\<not> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
proof -
assume 1: "\<not> is_Prim\<^sub>0 u"
have "Src v = Trg (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
using u v w I2 [of u] by (cases v, simp_all)
hence "Nml (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and>
Dom (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and>
Cod (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Cod v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and>
Src (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u \<and> Trg (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Trg v"
using u v w I1 [of "w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u"] I2 [of u] by argo
moreover have "v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = (v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u"
using 1 by (cases u, simp_all)
moreover have "(Dom v \<^bold>\<star> Dom w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u = Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> Cod w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u = Cod v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<langle>a\<^bold>\<rangle>\<^sub>0 = Trg u"
show "Nml (Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<and>
Dom (Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Dom (Trg u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u \<and>
Cod (Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Cod (Trg u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u \<and>
Src (Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u \<and> Trg (Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) " using assms 0 by blast
show "Dom (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Dom t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u" using assms 0 by blast
show "Cod (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Cod t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u" using assms 0 by blast
show "Src (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Src u" using assms 0 by blast
show "Trg (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<in> HHom (Src u) (Trg t)"
and "t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<in> VHom (Dom t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u) (Cod t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Dom t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Cod t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 "\<not> is_Prim\<^sub>0 t" and "\<not> is_Prim\<^sub>0 u"
shows "is_Hcomp (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
proof -
have "\<lbrakk> \<not> is_Hcomp (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u); Nml t; Nml u; Src t = Trg u; \<not> is_Prim\<^sub>0 t; \<not> is_Prim\<^sub>0 u \<rbrakk>
\<Longrightarrow> False"
proof (induct t, simp_all add: Nml_HcompD)
fix a
assume a: "obj a"
assume u: "Nml u"
assume 1: "\<not> is_Hcomp u"
assume 2: "\<not> is_Prim\<^sub>0 (Trg u)"
show "False"
using u 1 2 by (cases u; simp)
next
fix v w
assume I2: "\<not> is_Hcomp (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<Longrightarrow> False"
assume vw: "Nml (v \<^bold>\<star> w)"
assume u: "Nml u"
assume 1: "\<not> is_Hcomp ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
assume 2: "\<not> is_Prim\<^sub>0 u"
assume 3: "Src w = Trg u"
show "False"
proof -
have v: "v = \<^bold>\<langle>un_Prim v\<^bold>\<rangle>"
using vw Nml_HcompD by metis
have w: "Nml w \<and> \<not> is_Prim\<^sub>0 w \<and> \<^bold>\<langle>src (un_Prim v)\<^bold>\<rangle>\<^sub>0 = Trg w"
using vw Nml_HcompD [of v w] by blast
have "(v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = v \<^bold>\<star> (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
proof -
have "(v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = \<^bold>\<langle>un_Prim v\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
using u v 2 by (cases u, simp_all)
also have "... = \<^bold>\<langle>un_Prim v\<^bold>\<rangle> \<^bold>\<star> (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
using u w I2 by fastforce
also have "... = v \<^bold>\<star> (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 \<open>
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.
\<close>
primrec dim :: "'a term \<Rightarrow> nat"
where "dim \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = 0"
| "dim \<^bold>\<langle>\<mu>\<^bold>\<rangle> = 1"
| "dim (t \<^bold>\<star> u) = (dim t + dim u)"
| "dim (t \<^bold>\<cdot> u) = dim t"
| "dim \<^bold>\<l>\<^bold>[t\<^bold>] = dim t"
| "dim \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = dim t"
| "dim \<^bold>\<r>\<^bold>[t\<^bold>] = dim t"
| "dim \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = dim t"
| "dim \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = dim t + dim u + dim v"
| "dim \<^bold>\<a>\<^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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
proof -
have "\<And>n t u v. \<lbrakk> dim t = n; Nml t; Nml u; Nml v; Src t = Trg u; Src u = Trg v \<rbrakk> \<Longrightarrow>
(t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
proof -
fix n
show "\<And>t u v. \<lbrakk> dim t = n; Nml t; Nml u; Nml v; Src t = Trg u; Src u = Trg v \<rbrakk> \<Longrightarrow>
(t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
proof (induction n rule: nat_less_induct)
fix n :: nat and t :: "'a term" and u v
assume I: "\<forall>m<n. \<forall>t u v. dim t = m \<longrightarrow> Nml t \<longrightarrow> Nml u \<longrightarrow> Nml v \<longrightarrow>
Src t = Trg u \<longrightarrow> Src u = Trg v \<longrightarrow>
(t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
proof -
have "is_Prim\<^sub>0 t \<Longrightarrow> ?thesis" by (cases t; simp)
moreover have "\<not>is_Prim\<^sub>0 t \<Longrightarrow> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
by (cases t; simp; cases u; simp)
moreover have "\<not> is_Prim\<^sub>0 t \<Longrightarrow> \<not> is_Prim\<^sub>0 u \<Longrightarrow> is_Prim\<^sub>0 v \<Longrightarrow> ?thesis"
proof -
assume 1: "\<not> is_Prim\<^sub>0 t"
assume 2: "\<not> is_Prim\<^sub>0 u"
assume 3: "is_Prim\<^sub>0 v"
have "\<not>is_Prim\<^sub>0 (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u", simp_all)
qed
moreover have "\<not>is_Prim\<^sub>0 t \<and> \<not> is_Prim\<^sub>0 u \<and> \<not> is_Prim\<^sub>0 v \<and> is_Prim t \<Longrightarrow> ?thesis"
using v by (cases t, simp_all, cases u, simp_all; cases v, simp_all)
moreover have "\<not>is_Prim\<^sub>0 t \<and> \<not> is_Prim\<^sub>0 u \<and> \<not> is_Prim\<^sub>0 v \<and> is_Hcomp t \<Longrightarrow> ?thesis"
proof (cases t, simp_all)
fix w :: "'a term" and x :: "'a term"
assume 1: " \<not> is_Prim\<^sub>0 u \<and> \<not> is_Prim\<^sub>0 v"
assume 2: "t = (w \<^bold>\<star> x)"
show "((w \<^bold>\<star> x) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = (w \<^bold>\<star> x) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
proof -
have w: "w = \<^bold>\<langle>un_Prim w\<^bold>\<rangle>"
using t 1 2 Nml_HcompD by metis
have x: "Nml x"
using t w 1 2 by (metis Nml.simps(3))
have "((w \<^bold>\<star> x) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (x \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v"
using u v w x 1 2 by (cases u, simp_all)
also have "... = (w \<^bold>\<star> (x \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> ((x \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
using u v w x 1 by (cases u, simp_all; cases v, simp_all)
also have "... = w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (x \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v))"
proof -
have "dim x < dim t"
using 2 w by (cases w; simp)
moreover have "Src x = Trg u \<and> 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>\<star> x) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
proof -
have 3: "is_Hcomp (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v)"
using u v uv 1 is_Hcomp_HcompNml by auto
obtain u' :: "'a term" and v' where uv': "u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v = u' \<^bold>\<star> 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 \<or> is_Prim t \<or> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> t = t"
proof -
have "Nml t \<Longrightarrow> Trg t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> t = t"
proof (induct t, simp_all add: Nml_HcompD)
fix u v
assume uv: "Nml (u \<^bold>\<star> v)"
assume I1: "Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = u"
have 1: "Nml u \<and> Nml v \<and> Src u = Trg v"
using uv Nml_HcompD by blast
have 2: "Trg (u \<^bold>\<star> v) = Trg u"
using uv by auto
show "Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<star> v = u \<^bold>\<star> v"
proof -
have "Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<star> v = Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v"
using uv HcompNml_Nml by simp
also have "... = (Trg u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v"
using 1 2 HcompNml_assoc Src_Trg Nml_Trg Nml_implies_Arr by simp
also have "... = u \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src t = t"
proof -
have "Nml t \<Longrightarrow> t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src t = t"
proof (induct t, simp_all add: Nml_HcompD)
fix u v
assume uv: "Nml (u \<^bold>\<star> v)"
assume I2: "v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src v = v"
have 1: "Nml u \<and> Nml v \<and> Src u = Trg v"
using uv Nml_HcompD by blast
have 2: "Src (u \<^bold>\<star> v) = Src v"
using uv Trg_HcompNml by auto
show "(u \<^bold>\<star> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src v = u \<^bold>\<star> v"
proof -
have "(u \<^bold>\<star> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src v = (u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src v"
using uv HcompNml_Nml by simp
also have "... = u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Src v)"
using 1 2 HcompNml_assoc Trg_Src Nml_Src Nml_implies_Arr by simp
also have "... = u \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
proof -
have "\<And>u. \<lbrakk> Can t \<and> Nml t; Can u \<and> Nml u; Src t = Trg u \<rbrakk> \<Longrightarrow> Can (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
proof (induct t, simp_all add: HcompNml_Trg_Nml HcompNml_Nml_Src)
show "\<And>x u. ide x \<and> arr x \<Longrightarrow> Can u \<and> Nml u \<Longrightarrow> \<^bold>\<langle>src x\<^bold>\<rangle>\<^sub>0 = Trg u \<Longrightarrow> Can (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 "\<And>v w u.
(\<And>u. Nml v \<Longrightarrow> Can u \<and> Nml u \<Longrightarrow> Trg w = Trg u \<Longrightarrow> Can (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)) \<Longrightarrow>
(\<And>ua. Nml w \<Longrightarrow> Can ua \<and> Nml ua \<Longrightarrow> Trg u = Trg ua \<Longrightarrow> Can (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> ua)) \<Longrightarrow>
Can v \<and> Can w \<and> Src v = Trg w \<and> Nml (v \<^bold>\<star> w) \<Longrightarrow>
Can u \<and> Nml u \<Longrightarrow> Src w = Trg u \<Longrightarrow> Can ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Inv t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv u"
proof -
have "\<And>u. \<lbrakk> Can t \<and> Nml t; Can u \<and> Nml u; Src t = Trg u \<rbrakk>
\<Longrightarrow> Inv (t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Inv t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv u"
proof (induct t, simp_all add: HcompNml_Trg_Nml HcompNml_Nml_Src)
show "\<And>x u. \<lbrakk> obj x; Can u \<and> Nml u; \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 = Trg u \<rbrakk> \<Longrightarrow> Inv u = Inv (Trg u) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv u"
by (metis HcompNml.simps(1) Inv.simps(1))
show "\<And>x u. ide x \<and> arr x \<Longrightarrow> Can u \<and> Nml u \<Longrightarrow> \<^bold>\<langle>src x\<^bold>\<rangle>\<^sub>0 = Trg u \<Longrightarrow>
Inv (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = \<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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: "\<And>x. Nml v \<Longrightarrow> Can x \<and> Nml x \<Longrightarrow> Trg w = Trg x \<Longrightarrow>
Inv (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Inv v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv x"
assume I2: "\<And>x. Nml w \<Longrightarrow> Can x \<and> Nml x \<Longrightarrow> Trg u = Trg x \<Longrightarrow>
Inv (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x) = Inv w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv x"
assume vw: "Can v \<and> Can w \<and> Src v = Trg w \<and> Nml (v \<^bold>\<star> w)"
assume wu: "Src w = Trg u"
assume u: "Can u \<and> Nml u"
have v: "Can v \<and> Nml v"
using vw Nml_HcompD by blast
have w: "Can w \<and> Nml w"
using v vw by (cases v, simp_all)
show "Inv ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = (Inv v \<^bold>\<star> Inv w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv u"
proof -
have "is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
apply (cases u) by simp_all
moreover have "\<not> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
proof -
assume 1: "\<not> is_Prim\<^sub>0 u"
have "Inv ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) = Inv (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u))"
using 1 by (cases u, simp_all)
also have "... = Inv v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv (w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
using u v w vw wu I1 Nml_HcompNml Can_HcompNml Nml_Inv Can_Inv
by simp
also have "... = Inv v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (Inv w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv u)"
using u v w I2 Nml_HcompNml by simp
also have "... = (Inv v \<^bold>\<star> Inv w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 \<open>
The following function defines vertical composition for compatible normal terms,
by ``pushing the composition down'' to arrows of @{text V}.
\<close>
fun VcompNml :: "'a term \<Rightarrow> 'a term \<Rightarrow> 'a term" (infixr "\<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor>" 55)
where "\<^bold>\<langle>\<nu>\<^bold>\<rangle>\<^sub>0 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u = u"
| "\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<langle>\<mu>\<^bold>\<rangle> = \<^bold>\<langle>\<nu> \<cdot> \<mu>\<^bold>\<rangle>"
| "(u \<^bold>\<star> v) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (w \<^bold>\<star> x) = (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w \<^bold>\<star> v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x)"
| "t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = t"
| "t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> _ = undefined \<^bold>\<cdot> undefined"
text \<open>
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.
\<close>
lemma Nml_VcompNml:
assumes "Nml t" and "Nml u" and "Dom t = Cod u"
shows "Nml (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u)"
and "Dom (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u"
and "Cod (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Cod t"
proof -
have 0: "\<And>u. \<lbrakk> Nml t; Nml u; Dom t = Cod u \<rbrakk> \<Longrightarrow>
Nml (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<and> Dom (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u \<and> Cod (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Cod t"
proof (induct t, simp_all add: Nml_HcompD)
show "\<And>x u. obj x \<Longrightarrow> Nml u \<Longrightarrow> \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 = Cod u \<Longrightarrow>
Nml (Cod u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<and> Dom (Cod u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u \<and>
Cod (Cod u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Cod (Cod u)"
by (metis Cod.simps(1) VcompNml.simps(1))
fix \<nu> u
assume \<nu>: "arr \<nu>"
assume u: "Nml u"
assume 1: "\<^bold>\<langle>dom \<nu>\<^bold>\<rangle> = Cod u"
show "Nml (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<and> Dom (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u \<and> Cod (\<^bold>\<langle>\<nu>\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = \<^bold>\<langle>cod \<nu>\<^bold>\<rangle>"
using \<nu> u 1 by (cases u, simp_all)
next
fix u v w
assume I2: "\<And>u. \<lbrakk> Nml u; Dom w = Cod u \<rbrakk> \<Longrightarrow>
Nml (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<and> Dom (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u \<and> Cod (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Cod w"
assume vw: "Nml (v \<^bold>\<star> w)"
have v: "Nml v"
using vw Nml_HcompD by blast
have w: "Nml w"
using vw Nml_HcompD by blast
assume u: "Nml u"
assume 1: "(Dom v \<^bold>\<star> Dom w) = Cod u"
show "Nml ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<and> Dom ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u \<and>
Cod ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Cod v \<^bold>\<star> Cod w"
using u v w 1
proof (cases u, simp_all)
fix x y
assume 2: "u = x \<^bold>\<star> y"
have 4: "is_Prim x \<and> x = \<^bold>\<langle>un_Prim x\<^bold>\<rangle> \<and> arr (un_Prim x) \<and> Nml y \<and> \<not> is_Prim\<^sub>0 y"
using u 2 by (cases x, cases y, simp_all)
have 5: "is_Prim v \<and> v = \<^bold>\<langle>un_Prim v\<^bold>\<rangle> \<and> arr (un_Prim v) \<and> Nml w \<and> \<not> is_Prim\<^sub>0 w"
using v w vw by (cases v, simp_all)
have 6: "dom (un_Prim v) = cod (un_Prim x) \<and> Dom w = Cod y"
proof -
have "\<^bold>\<langle>src (un_Prim v)\<^bold>\<rangle>\<^sub>0 = Trg w" using vw Nml_HcompD [of v w] by simp
thus "dom (un_Prim v) = cod (un_Prim x) \<and> 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>\<star> w) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u = \<^bold>\<langle>un_Prim v \<cdot> un_Prim x\<^bold>\<rangle> \<^bold>\<star> w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y"
using 2 4 5 6 VcompNml.simps(2) [of "un_Prim v" "un_Prim x"] by simp
moreover have "Nml (\<^bold>\<langle>un_Prim v \<cdot> un_Prim x\<^bold>\<rangle> \<^bold>\<star> w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y)"
proof -
have "Nml (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y)"
using I2 4 5 6 by simp
moreover have "\<not> is_Prim\<^sub>0 (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<langle>src (un_Prim v \<cdot> un_Prim x)\<^bold>\<rangle>\<^sub>0 = Trg (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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 \<cdot> un_Prim x" "(w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y)"]
by simp
qed
ultimately show "Nml (v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x \<^bold>\<star> w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y) \<and>
Dom (v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x) = Dom x \<and> Dom (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y) = Dom y \<and>
Cod (v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x) = Cod v \<and> Cod (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u)" using assms 0 by blast
show "Dom (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Dom u" using assms 0 by blast
show "Cod (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u \<in> HHom (Src u) (Trg u)" and "t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u \<in> VHom (Dom u) (Cod t)"
proof -
show 1: "t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u \<in> VHom (Dom u) (Cod t)"
using assms Nml_VcompNml Nml_implies_Arr by simp
show "t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u \<in> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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 \<Longrightarrow> Cod t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (Dom t) = t"
proof -
have "Nml t \<Longrightarrow> t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
proof -
have "\<And>u v. \<lbrakk> Nml t; Nml u; Nml v; Dom t = Cod u; Dom u = Cod v \<rbrakk> \<Longrightarrow>
(t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
proof (induct t, simp_all)
show "\<And>x u v. obj x \<Longrightarrow> Nml u \<Longrightarrow> Nml v \<Longrightarrow> \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 = Cod u \<Longrightarrow> Dom u = Cod v \<Longrightarrow>
u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = Cod u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<langle>dom f\<^bold>\<rangle> = Cod u"
assume 2: "Dom u = Cod v"
show "(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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: "\<And>u v. \<lbrakk> Nml w; Nml u; Nml v; Dom w = Cod u; Dom u = Cod v \<rbrakk> \<Longrightarrow>
(w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
assume I2: "\<And>u v. \<lbrakk> Nml x; Nml u; Nml v; Dom x = Cod u; Dom u = Cod v \<rbrakk> \<Longrightarrow>
(x \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = x \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
assume wx: "Nml (w \<^bold>\<star> x)"
assume u: "Nml u"
assume v: "Nml v"
assume 1: "(Dom w \<^bold>\<star> Dom x) = Cod u"
assume 2: "Dom u = Cod v"
show "((w \<^bold>\<star> x) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v = (w \<^bold>\<star> x) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<star> u2)"
assume 4: "v = (v1 \<^bold>\<star> v2)"
show "(w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u1) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1 = w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1 \<and>
(x \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u2) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2 = x \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2"
proof -
have "Nml u1 \<and> Nml u2"
using u 3 Nml_HcompD by blast
moreover have "Nml v1 \<and> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u)"
proof -
have "\<And>u. \<lbrakk> Ide t; Ide u; Nml t; Nml u; Dom t = Cod u \<rbrakk> \<Longrightarrow> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u)"
proof -
have "\<And>u. \<lbrakk> Can t \<and> Nml t; Can u \<and> Nml u; Dom t = Cod u \<rbrakk> \<Longrightarrow> Can (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u)"
proof (induct t, simp_all)
fix t u v
assume I1: "\<And>v. \<lbrakk> Nml t; Can v \<and> Nml v; Dom t = Cod v \<rbrakk> \<Longrightarrow> Can (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
assume I2: "\<And>v. \<lbrakk> Nml u; Can v \<and> Nml v; Dom u = Cod v \<rbrakk> \<Longrightarrow> Can (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
assume tu: "Can t \<and> Can u \<and> Src t = Trg u \<and> Nml (t \<^bold>\<star> u)"
have t: "Can t \<and> Nml t"
using tu Nml_HcompD by blast
have u: "Can u \<and> Nml u"
using tu Nml_HcompD by blast
assume v: "Can v \<and> Nml v"
assume 1: "(Dom t \<^bold>\<star> Dom u) = Cod v"
show "Can ((t \<^bold>\<star> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v)"
proof -
have 2: "(Dom t \<^bold>\<star> 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>\<star> x)"
have "Can w \<and> Nml w" using v wx Nml_HcompD Can.simps(3) by blast
moreover have "Can x \<and> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w) \<and> Can (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x) \<and> Src (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w) = Trg (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv t"
proof -
have "\<And>u. \<lbrakk> Can t \<and> Nml t; Can u \<and> Nml u; Dom t = Cod u \<rbrakk> \<Longrightarrow>
Inv (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv t"
proof (induct t, simp_all)
show "\<And>x u. \<lbrakk> obj x; Can u \<and> Nml u; \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 = Cod u \<rbrakk> \<Longrightarrow> Inv u = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv (Cod u)"
by (simp add: Can_implies_Arr)
show "\<And>x u. \<lbrakk> ide x \<and> arr x; Can u \<and> Nml u; \<^bold>\<langle>x\<^bold>\<rangle> = Cod u \<rbrakk> \<Longrightarrow>
Inv u = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv (Cod u)"
by (simp add: Can_implies_Arr)
fix v w u
assume vw: "Can v \<and> Can w \<and> Src v = Trg w \<and> Nml (v \<^bold>\<star> w)"
have v: "Can v \<and> Nml w"
using vw Nml_HcompD by blast
have w: "Can w \<and> Nml w"
using vw Nml_HcompD by blast
assume I1: "\<And>x. \<lbrakk> Nml v; Can x \<and> Nml x; Dom v = Cod x \<rbrakk> \<Longrightarrow>
Inv (v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x) = Inv x \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv v"
assume I2: "\<And>x. \<lbrakk> Nml w; Can x \<and> Nml x; Dom w = Cod x \<rbrakk> \<Longrightarrow>
Inv (w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x) = Inv x \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv w"
assume u: "Can u \<and> Nml u"
assume 1: "(Dom v \<^bold>\<star> Dom w) = Cod u"
show "Inv ((v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (Inv v \<^bold>\<star> Inv w)"
using v 1
proof (cases w, simp_all)
show "\<And>\<mu>. obj \<mu> \<Longrightarrow> Dom v \<^bold>\<star> \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 = Cod u \<Longrightarrow> w = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0 \<Longrightarrow> Can v \<Longrightarrow>
Inv ((v \<^bold>\<star> \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (Inv v \<^bold>\<star> \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0)"
using Nml_HcompD(5) is_Prim\<^sub>0_def vw by blast
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Dom v \<^bold>\<star> \<^bold>\<langle>dom \<mu>\<^bold>\<rangle> = Cod u \<Longrightarrow> w = \<^bold>\<langle>\<mu>\<^bold>\<rangle> \<Longrightarrow> Can v \<Longrightarrow>
Inv ((v \<^bold>\<star> \<^bold>\<langle>\<mu>\<^bold>\<rangle>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (Inv v \<^bold>\<star> \<^bold>\<langle>inv \<mu>\<^bold>\<rangle>)"
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 "\<And>y z. Nml (y \<^bold>\<star> z) \<Longrightarrow> Dom v \<^bold>\<star> Dom y \<^bold>\<star> Dom z = Cod u \<Longrightarrow>
w = y \<^bold>\<star> z \<Longrightarrow> Can v \<Longrightarrow>
Inv ((v \<^bold>\<star> y \<^bold>\<star> z) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (Inv v \<^bold>\<star> Inv y \<^bold>\<star> Inv z)"
proof -
fix y z
assume 2: "Nml (y \<^bold>\<star> z)"
assume yz: "w = y \<^bold>\<star> z"
show "Inv ((v \<^bold>\<star> y \<^bold>\<star> z) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u) = Inv u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (Inv v \<^bold>\<star> Inv y \<^bold>\<star> 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 "\<lbrakk> Can t; Nml t \<rbrakk> \<Longrightarrow> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> t = Dom t"
using assms Can_and_Nml_implies_Ide Ide_in_Hom by simp
text \<open>
The next fact is a syntactic version of the interchange law, for normal terms.
\<close>
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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)"
proof -
have "\<And>u v w. \<lbrakk> Nml t; Nml u; Nml v; Nml w; VSeq t v; VSeq u w;
Src t = Trg u; Src v = Trg w \<rbrakk> \<Longrightarrow>
(t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = (t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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 "\<And>x. Arr v \<and> \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 = Cod v \<Longrightarrow> (Cod v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w"
using u v w uw by (cases v) simp_all
show "\<And>x. \<lbrakk> arr x; Arr v \<and> \<^bold>\<langle>dom x\<^bold>\<rangle> = Cod v; \<^bold>\<langle>src x\<^bold>\<rangle>\<^sub>0 = Trg u; Src v = Trg w \<rbrakk> \<Longrightarrow>
(\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = \<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w"
proof -
fix x
assume x: "arr x"
assume 1: "Arr v \<and> \<^bold>\<langle>dom x\<^bold>\<rangle> = Cod v"
assume tu: "\<^bold>\<langle>src x\<^bold>\<rangle>\<^sub>0 = Trg u"
assume vw: "Src v = Trg w"
show "(\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = \<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w"
proof -
have 2: "v = \<^bold>\<langle>un_Prim v\<^bold>\<rangle> \<and> arr (un_Prim v)" using v 1 by (cases v) simp_all
have "is_Prim\<^sub>0 u \<Longrightarrow> ?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 "\<not> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
proof -
assume 3: "\<not> is_Prim\<^sub>0 u"
hence 4: "\<not> is_Prim\<^sub>0 w" using u w uw by (cases u, simp_all; cases w, simp_all)
have "(\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<star> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<star> w)"
proof -
have "\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = \<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<star> u"
using u x 3 HcompNml_Nml by (cases u, simp_all)
moreover have "v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w = v \<^bold>\<star> 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>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v) \<^bold>\<star> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)" by simp
also have "... = (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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>\<star> t2)"
assume I1: "\<And>u v w. \<lbrakk> Nml t1; Nml u; Nml v; Nml w;
Arr v \<and> Dom t1 = Cod v; VSeq u w;
Trg t2 = Trg u; Src v = Trg w \<rbrakk> \<Longrightarrow>
(t1 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = t1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w"
assume I2: "\<And>u' v w. \<lbrakk> Nml t2; Nml u'; Nml v; Nml w;
Arr v \<and> Dom t2 = Cod v; VSeq u' w;
Trg u = Trg u'; Src v = Trg w \<rbrakk> \<Longrightarrow>
(t2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u') \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = (t2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u' \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)"
have t1: "t1 = \<^bold>\<langle>un_Prim t1\<^bold>\<rangle> \<and> arr (un_Prim t1) \<and> Nml t1"
using t12 by (cases t1, simp_all)
have t2: "Nml t2 \<and> \<not>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 \<and> Arr t2 \<and> Src t1 = Trg t2 \<and> Arr v \<and> (Dom t1 \<^bold>\<star> Dom t2) = Cod v"
show "((t1 \<^bold>\<star> t2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = (t1 \<^bold>\<star> t2) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w"
proof -
have "is_Prim\<^sub>0 u \<Longrightarrow> ?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 "\<not> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
proof -
assume u': "\<not> is_Prim\<^sub>0 u"
hence w': "\<not> 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>\<star> v2"
have v1: "v1 = \<^bold>\<langle>un_Prim v1\<^bold>\<rangle> \<and> arr (un_Prim v1) \<and> Nml v1"
using v v12 by (cases v1, simp_all)
have v2: "Nml v2 \<and> \<not> is_Prim\<^sub>0 v2"
using v v12 by (cases v1, simp_all)
have 2: "v = (\<^bold>\<langle>un_Prim v1\<^bold>\<rangle> \<^bold>\<star> v2)"
using v1 v12 by simp
show "((t1 \<^bold>\<star> t2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> ((v1 \<^bold>\<star> v2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) = (t1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1 \<^bold>\<star> t2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w"
proof -
have 3: "(t1 \<^bold>\<star> t2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u = t1 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (t2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
using u u' by (cases u, simp_all)
have 4: "v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w = v1 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> t2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> ((v1 \<^bold>\<star> v2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w)
= (t1 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (t2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v1 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (v2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w))"
using 3 4 v12 by simp
also have "... = (t1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> ((t2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (v2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w))"
proof -
have "is_Hcomp (t2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u)"
using t2 u u' tu is_Hcomp_HcompNml by auto
moreover have "is_Hcomp (v2 \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (t2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)"
using u w tu uw vw t2 v2 1 2 Nml_implies_Arr I2 by auto
also have "... = ((t1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1) \<^bold>\<star> (t2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)"
proof -
have "\<not>is_Prim\<^sub>0 (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)"
using u w u' w' by (cases u, simp_all; cases w, simp_all)
hence "((t1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1) \<^bold>\<star> (t2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w)
= (t1 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v1) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> ((t2 \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> v2) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> w))"
by (cases "u \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> 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 \<open>
The following function reduces a formal arrow to normal form.
\<close>
fun Nmlize :: "'a term \<Rightarrow> 'a term" ("\<^bold>\<lfloor>_\<^bold>\<rfloor>")
where "\<^bold>\<lfloor>\<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0\<^bold>\<rfloor> = \<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^sub>0"
| "\<^bold>\<lfloor>\<^bold>\<langle>\<mu>\<^bold>\<rangle>\<^bold>\<rfloor> = \<^bold>\<langle>\<mu>\<^bold>\<rangle>"
| "\<^bold>\<lfloor>t \<^bold>\<star> u\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>t \<^bold>\<cdot> u\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>\<^bold>\<l>\<^bold>[t\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>\<^bold>\<r>\<^bold>[t\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<^bold>\<rfloor> = (\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>"
| "\<^bold>\<lfloor>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>)"
lemma Nml_Nmlize:
assumes "Arr t"
shows "Nml \<^bold>\<lfloor>t\<^bold>\<rfloor>" and "Src \<^bold>\<lfloor>t\<^bold>\<rfloor> = Src t" and "Trg \<^bold>\<lfloor>t\<^bold>\<rfloor> = Trg t"
and "Dom \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t\<^bold>\<rfloor>" and "Cod \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>"
proof -
have 0: "Arr t \<Longrightarrow> Nml \<^bold>\<lfloor>t\<^bold>\<rfloor> \<and> Src \<^bold>\<lfloor>t\<^bold>\<rfloor> = Src t \<and> Trg \<^bold>\<lfloor>t\<^bold>\<rfloor> = Trg t \<and>
Dom \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<and> Cod \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>"
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>\<lfloor>t\<^bold>\<rfloor>"
assume 3: "Src \<^bold>\<lfloor>t\<^bold>\<rfloor> = Src t"
assume 4: "Trg \<^bold>\<lfloor>t\<^bold>\<rfloor> = Trg t"
assume 5: "Dom \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t\<^bold>\<rfloor>"
assume 6: "Cod \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>"
have 7: "Nml \<^bold>\<lfloor>Dom t\<^bold>\<rfloor>"
using 2 5 Nml_Dom by fastforce
have 8: "Trg \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Trg t\<^bold>\<rfloor>"
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>\<lfloor>Cod t\<^bold>\<rfloor>"
using 2 6 Nml_Cod by fastforce
have 10: "Src \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Src t\<^bold>\<rfloor>"
using 1 2 3 Nml_Src Obj_Src
by (metis Nml.elims(2) Nmlize.simps(1) Nmlize.simps(2) Obj.simps(3))
show "\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>Trg t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t\<^bold>\<rfloor>"
using 2 5 7 8 Nml_implies_Arr Trg_Dom HcompNml_Trg_Nml by metis
show "\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> = \<^bold>\<lfloor>Trg t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>"
using 2 6 8 9 Nml_implies_Arr Trg_Cod HcompNml_Trg_Nml by metis
show "\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>"
using 2 5 7 10 Nml_implies_Arr Src_Dom HcompNml_Nml_Src by metis
show "\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>"
using 2 6 9 10 Nml_implies_Arr Src_Cod HcompNml_Nml_Src by metis
next
fix t1 t2 t3
assume "Nml \<^bold>\<lfloor>t1\<^bold>\<rfloor>" and "Nml \<^bold>\<lfloor>t2\<^bold>\<rfloor>" and "Nml \<^bold>\<lfloor>t3\<^bold>\<rfloor>"
assume "Src t1 = Trg t2" and "Src t2 = Trg t3"
assume "Src \<^bold>\<lfloor>t1\<^bold>\<rfloor> = Trg t2" and "Src \<^bold>\<lfloor>t2\<^bold>\<rfloor> = Trg t3"
assume "Trg \<^bold>\<lfloor>t1\<^bold>\<rfloor> = Trg t1" and "Trg \<^bold>\<lfloor>t2\<^bold>\<rfloor> = Trg t2" and "Trg \<^bold>\<lfloor>t3\<^bold>\<rfloor> = Trg t3"
assume "Dom \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t1\<^bold>\<rfloor>" and "Cod \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t1\<^bold>\<rfloor>" and "Dom \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>"
and "Cod \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor>" and "Dom \<^bold>\<lfloor>t3\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor>" and "Cod \<^bold>\<lfloor>t3\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t3\<^bold>\<rfloor>"
show "\<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor> =
(\<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor>"
using Nml_Dom Nml_implies_Arr Src_Dom Trg_Dom
HcompNml_assoc [of "\<^bold>\<lfloor>Dom t1\<^bold>\<rfloor>" "\<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>" "\<^bold>\<lfloor>Dom t3\<^bold>\<rfloor>"]
\<open>Nml \<^bold>\<lfloor>t1\<^bold>\<rfloor>\<close> \<open>Nml \<^bold>\<lfloor>t2\<^bold>\<rfloor>\<close> \<open>Nml \<^bold>\<lfloor>t3\<^bold>\<rfloor>\<close>
\<open>Dom \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t1\<^bold>\<rfloor>\<close> \<open>Dom \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>\<close> \<open>Dom \<^bold>\<lfloor>t3\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor>\<close>
\<open>Src \<^bold>\<lfloor>t1\<^bold>\<rfloor> = Trg t2\<close> \<open>Trg \<^bold>\<lfloor>t2\<^bold>\<rfloor> = Trg t2\<close>
\<open>Src \<^bold>\<lfloor>t2\<^bold>\<rfloor> = Trg t3\<close> \<open>Trg \<^bold>\<lfloor>t3\<^bold>\<rfloor> = Trg t3\<close>
by metis
show "\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t3\<^bold>\<rfloor> = (\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t3\<^bold>\<rfloor>"
using Nml_Cod Nml_implies_Arr Src_Cod Trg_Cod
HcompNml_assoc [of "\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor>" "\<^bold>\<lfloor>Cod t2\<^bold>\<rfloor>" "\<^bold>\<lfloor>Cod t3\<^bold>\<rfloor>"]
\<open>Nml \<^bold>\<lfloor>t1\<^bold>\<rfloor>\<close> \<open>Nml \<^bold>\<lfloor>t2\<^bold>\<rfloor>\<close> \<open>Nml \<^bold>\<lfloor>t3\<^bold>\<rfloor>\<close>
\<open>Cod \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t1\<^bold>\<rfloor>\<close> \<open>Cod \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor>\<close> \<open>Cod \<^bold>\<lfloor>t3\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t3\<^bold>\<rfloor>\<close>
\<open>Src \<^bold>\<lfloor>t1\<^bold>\<rfloor> = Trg t2\<close> \<open>Trg \<^bold>\<lfloor>t2\<^bold>\<rfloor> = Trg t2\<close>
\<open>Src \<^bold>\<lfloor>t2\<^bold>\<rfloor> = Trg t3\<close> \<open>Trg \<^bold>\<lfloor>t3\<^bold>\<rfloor> = Trg t3\<close>
by metis
qed
show "Nml \<^bold>\<lfloor>t\<^bold>\<rfloor>" using assms 0 by blast
show "Src \<^bold>\<lfloor>t\<^bold>\<rfloor> = Src t" using assms 0 by blast
show "Trg \<^bold>\<lfloor>t\<^bold>\<rfloor> = Trg t" using assms 0 by blast
show "Dom \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t\<^bold>\<rfloor>" using assms 0 by blast
show "Cod \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>" using assms 0 by blast
qed
lemma Nmlize_in_Hom [intro]:
assumes "Arr t"
shows "\<^bold>\<lfloor>t\<^bold>\<rfloor> \<in> HHom (Src t) (Trg t)" and "\<^bold>\<lfloor>t\<^bold>\<rfloor> \<in> VHom \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>"
using assms Nml_Nmlize Nml_implies_Arr by auto
lemma Nmlize_Src:
assumes "Arr t"
shows "\<^bold>\<lfloor>Src t\<^bold>\<rfloor> = Src \<^bold>\<lfloor>t\<^bold>\<rfloor>" and "\<^bold>\<lfloor>Src t\<^bold>\<rfloor> = Src t"
proof -
have 1: "Obj (Src t)"
using assms by simp
obtain a where a: "obj a \<and> Src t = \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0"
using 1 by (cases "Src t", simp_all)
show "\<^bold>\<lfloor>Src t\<^bold>\<rfloor> = Src t"
using a by simp
thus "\<^bold>\<lfloor>Src t\<^bold>\<rfloor> = Src \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Nmlize_in_Hom by simp
qed
lemma Nmlize_Trg:
assumes "Arr t"
shows "\<^bold>\<lfloor>Trg t\<^bold>\<rfloor> = Trg \<^bold>\<lfloor>t\<^bold>\<rfloor>" and "\<^bold>\<lfloor>Trg t\<^bold>\<rfloor> = Trg t"
proof -
have 1: "Obj (Trg t)"
using assms by simp
obtain a where a: "obj a \<and> Trg t = \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0"
using 1 by (cases "Trg t", simp_all)
show "\<^bold>\<lfloor>Trg t\<^bold>\<rfloor> = Trg t"
using a by simp
thus "\<^bold>\<lfloor>Trg t\<^bold>\<rfloor> = Trg \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Nmlize_in_Hom by simp
qed
lemma Nmlize_Dom:
assumes "Arr t"
shows "\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> = Dom \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Nmlize_in_Hom by simp
lemma Nmlize_Cod:
assumes "Arr t"
shows "\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> = Cod \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Nmlize_in_Hom by simp
lemma Ide_Nmlize_Ide:
assumes "Ide t"
shows "Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof -
have "Ide t \<Longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>"
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>\<lfloor>t\<^bold>\<rfloor>"
proof -
have "Can t \<Longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>"
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>\<lfloor>t\<^bold>\<rfloor>"
using assms Ide_Nmlize_Can Ide_implies_Can by auto
lemma Nmlize_Nml [simp]:
assumes "Nml t"
shows "\<^bold>\<lfloor>t\<^bold>\<rfloor> = t"
proof -
have "Nml t \<Longrightarrow> \<^bold>\<lfloor>t\<^bold>\<rfloor> = 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>\<lfloor>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Nml_Nmlize Nmlize_Nml by blast
lemma Nmlize_Hcomp:
assumes "Arr t" and "Arr u"
shows "\<^bold>\<lfloor>t \<^bold>\<star> u\<^bold>\<rfloor> = \<^bold>\<lfloor>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>u\<^bold>\<rfloor>\<^bold>\<rfloor>"
using assms Nmlize_Nmlize by simp
lemma Nmlize_Hcomp_Obj_Arr [simp]:
assumes "Arr u"
shows "\<^bold>\<lfloor>\<^bold>\<langle>b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> u\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
using assms by simp
lemma Nmlize_Hcomp_Arr_Obj [simp]:
assumes "Arr t" and "Src t = \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0"
shows "\<^bold>\<lfloor>t \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms HcompNml_Nml_Src Nmlize_in_Hom by simp
lemma Nmlize_Hcomp_Prim_Arr [simp]:
assumes "Arr u" and "\<not> is_Prim\<^sub>0 \<^bold>\<lfloor>u\<^bold>\<rfloor>"
shows "\<^bold>\<lfloor>\<^bold>\<langle>\<mu>\<^bold>\<rangle> \<^bold>\<star> u\<^bold>\<rfloor> = \<^bold>\<langle>\<mu>\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<lfloor>u\<^bold>\<rfloor>"
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>\<lfloor>(t \<^bold>\<star> u) \<^bold>\<star> v\<^bold>\<rfloor> = \<^bold>\<lfloor>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<star> (\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>v\<^bold>\<rfloor>)\<^bold>\<rfloor>"
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>\<lfloor>t \<^bold>\<star> u \<^bold>\<star> v\<^bold>\<rfloor> = \<^bold>\<lfloor>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>v\<^bold>\<rfloor>\<^bold>\<rfloor>"
using assms Nml_Nmlize Nmlize_Nmlize by (simp add: HcompNml_assoc)
lemma Nmlize_Vcomp_Cod_Arr:
assumes "Arr t"
shows "\<^bold>\<lfloor>Cod t \<^bold>\<cdot> t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof -
have "Arr t \<Longrightarrow> \<^bold>\<lfloor>Cod t \<^bold>\<cdot> t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof (induct t, simp_all)
show "\<And>x. arr x \<Longrightarrow> cod x \<cdot> x = x"
using comp_cod_arr by blast
fix t1 t2
show "\<And>t1 t2. \<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow> HSeq t1 t2 \<Longrightarrow>
(\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>"
using VcompNml_HcompNml Ide_Cod Nml_Nmlize Nmlize_in_Hom
by simp
show "\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow> VSeq t1 t2 \<Longrightarrow>
\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>"
using VcompNml_assoc [of "\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor>" "\<^bold>\<lfloor>t1\<^bold>\<rfloor>" "\<^bold>\<lfloor>t2\<^bold>\<rfloor>"] Ide_Cod
Nml_Nmlize
by simp
next
show "\<And>t. \<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<Longrightarrow> Arr t \<Longrightarrow> (\<^bold>\<lfloor>Trg t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
by (metis Arr.simps(6) Cod.simps(6) Nmlize.simps(3) Nmlize.simps(6)
Nmlize_Cod)
show "\<And>t. \<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<Longrightarrow> Arr t \<Longrightarrow> (\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
by (simp add: Nml_Nmlize(1) Nml_Nmlize(2) Nmlize_Src(2) HcompNml_Nml_Obj)
show "\<And>t1 t2 t3. \<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow>
\<^bold>\<lfloor>Cod t3\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor> = \<^bold>\<lfloor>t3\<^bold>\<rfloor> \<Longrightarrow>
Arr t1 \<and> Arr t2 \<and> Arr t3 \<and> Src t1 = Trg t2 \<and> Src t2 = Trg t3 \<Longrightarrow>
(\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t3\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> ((\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>) =
(\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>"
using VcompNml_HcompNml Ide_Cod HcompNml_in_Hom Nml_HcompNml
Nml_Nmlize Nmlize_in_Hom HcompNml_assoc
by simp
show "\<And>t1 t2 t3. \<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow>
\<^bold>\<lfloor>Cod t3\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor> = \<^bold>\<lfloor>t3\<^bold>\<rfloor> \<Longrightarrow>
Arr t1 \<and> Arr t2 \<and> Arr t3 \<and> Src t1 = Trg t2 \<and> Src t2 = Trg t3 \<Longrightarrow>
((\<^bold>\<lfloor>Cod t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t3\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>) =
\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>"
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>\<lfloor>t \<^bold>\<cdot> Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof -
have "Arr t \<Longrightarrow> \<^bold>\<lfloor>t \<^bold>\<cdot> Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof (induct t, simp_all)
show "\<And>x. arr x \<Longrightarrow> x \<cdot> local.dom x = x"
using comp_arr_dom by blast
fix t1 t2
show "\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow> HSeq t1 t2 \<Longrightarrow>
(\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>) = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>"
using VcompNml_HcompNml Ide_Dom HcompNml_in_Hom Nml_HcompNml
Nml_Nmlize Nmlize_in_Hom HcompNml_assoc
by simp
show "\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow> VSeq t1 t2 \<Longrightarrow>
(\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>"
using VcompNml_assoc [of "\<^bold>\<lfloor>t1\<^bold>\<rfloor>" "\<^bold>\<lfloor>t2\<^bold>\<rfloor>" "\<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>"] Ide_Dom Nml_Nmlize
by simp
next
show "\<And>t. \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<Longrightarrow> Arr t \<Longrightarrow> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>Trg t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t\<^bold>\<rfloor>) = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
by (simp add: Nml_Nmlize(1) Nml_Nmlize(3) Nmlize_Trg(2)
HcompNml_Obj_Nml bicategorical_language.Ide_Dom
bicategorical_language_axioms)
show "\<And>t. \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor> \<Longrightarrow> Arr t \<Longrightarrow> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>) = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
by (simp add: Nml_Nmlize(1) Nml_Nmlize(2) Nmlize_Src(2) HcompNml_Nml_Obj)
show "\<And>t1 t2 t3. \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow>
\<^bold>\<lfloor>t3\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor> = \<^bold>\<lfloor>t3\<^bold>\<rfloor> \<Longrightarrow>
Arr t1 \<and> Arr t2 \<and> Arr t3 \<and> Src t1 = Trg t2 \<and> Src t2 = Trg t3 \<Longrightarrow>
((\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> ((\<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor>) =
(\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>"
using VcompNml_HcompNml Ide_Dom HcompNml_in_Hom Nml_HcompNml
Nml_Nmlize Nmlize_in_Hom HcompNml_assoc
by simp
show "\<And>t1 t2 t3. \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> = \<^bold>\<lfloor>t1\<^bold>\<rfloor> \<Longrightarrow> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> = \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<Longrightarrow>
\<^bold>\<lfloor>t3\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor> = \<^bold>\<lfloor>t3\<^bold>\<rfloor> \<Longrightarrow>
Arr t1 \<and> Arr t2 \<and> Arr t3 \<and> Src t1 = Trg t2 \<and> Src t2 = Trg t3 \<Longrightarrow>
(\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>Dom t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t3\<^bold>\<rfloor>) =
\<^bold>\<lfloor>t1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>t3\<^bold>\<rfloor>"
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>\<lfloor>Inv t\<^bold>\<rfloor> = Inv \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof -
have "Can t \<Longrightarrow> \<^bold>\<lfloor>Inv t\<^bold>\<rfloor> = Inv \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof (induct t, simp_all)
fix u v
assume I1: "\<^bold>\<lfloor>Inv u\<^bold>\<rfloor> = Inv \<^bold>\<lfloor>u\<^bold>\<rfloor>"
assume I2: "\<^bold>\<lfloor>Inv v\<^bold>\<rfloor> = Inv \<^bold>\<lfloor>v\<^bold>\<rfloor>"
show "Can u \<and> Can v \<and> Src u = Trg v \<Longrightarrow> Inv \<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv \<^bold>\<lfloor>v\<^bold>\<rfloor> = Inv (\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>)"
using Inv_HcompNml Nml_Nmlize Can_implies_Arr Can_Nmlize_Can
I1 I2
by simp
show "Can u \<and> Can v \<and> Dom u = Cod v \<Longrightarrow> Inv \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> Inv \<^bold>\<lfloor>u\<^bold>\<rfloor> = Inv (\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>)"
using Inv_VcompNml Nml_Nmlize Can_implies_Arr Nmlize_in_Hom Can_Nmlize_Can
I1 I2
by simp
fix w
assume I3: "\<^bold>\<lfloor>Inv w\<^bold>\<rfloor> = Inv \<^bold>\<lfloor>w\<^bold>\<rfloor>"
assume uvw: "Can u \<and> Can v \<and> Can w \<and> Src u = Trg v \<and> Src v = Trg w"
show "Inv \<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (Inv \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv \<^bold>\<lfloor>w\<^bold>\<rfloor>) = Inv ((\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>)"
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>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv \<^bold>\<lfloor>v\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Inv \<^bold>\<lfloor>w\<^bold>\<rfloor> = Inv (\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>))"
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 \<open>
Function \<open>red\<close> defined below takes a formal identity @{term t} to a canonical arrow
\<open>f\<^bold>\<down> \<in> Hom f \<^bold>\<lfloor>f\<^bold>\<rfloor>\<close>. The auxiliary function \<open>red2\<close> takes a pair @{term "(f, g)"}
of normalized formal identities and produces a canonical arrow
\<open>f \<^bold>\<Down> g \<in> Hom (f \<^bold>\<star> g) \<^bold>\<lfloor>f \<^bold>\<star> g\<^bold>\<rfloor>\<close>.
\<close>
fun red2 (infixr "\<^bold>\<Down>" 53)
where "\<^bold>\<langle>b\<^bold>\<rangle>\<^sub>0 \<^bold>\<Down> u = \<^bold>\<l>\<^bold>[u\<^bold>]"
| "\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 = \<^bold>\<r>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>\<^bold>]"
| "\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> u = \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> u"
| "(t \<^bold>\<star> u) \<^bold>\<Down> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 = \<^bold>\<r>\<^bold>[t \<^bold>\<star> u\<^bold>]"
| "(t \<^bold>\<star> u) \<^bold>\<Down> v = (t \<^bold>\<Down> \<^bold>\<lfloor>u \<^bold>\<star> v\<^bold>\<rfloor>) \<^bold>\<cdot> (t \<^bold>\<star> (u \<^bold>\<Down> v)) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[t, u, v\<^bold>]"
| "t \<^bold>\<Down> u = undefined"
fun red ("_\<^bold>\<down>" [56] 56)
where "\<^bold>\<langle>f\<^bold>\<rangle>\<^sub>0\<^bold>\<down> = \<^bold>\<langle>f\<^bold>\<rangle>\<^sub>0"
| "\<^bold>\<langle>f\<^bold>\<rangle>\<^bold>\<down> = \<^bold>\<langle>f\<^bold>\<rangle>"
| "(t \<^bold>\<star> u)\<^bold>\<down> = (if Nml (t \<^bold>\<star> u) then t \<^bold>\<star> u else (\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>u\<^bold>\<rfloor>) \<^bold>\<cdot> (t\<^bold>\<down> \<^bold>\<star> u\<^bold>\<down>))"
| "t\<^bold>\<down> = undefined"
lemma red_Nml [simp]:
assumes "Nml a"
shows "a\<^bold>\<down> = a"
using assms by (cases a, simp_all)
lemma red2_Nml:
assumes "Nml (a \<^bold>\<star> b)"
shows "a \<^bold>\<Down> b = a \<^bold>\<star> b"
proof -
have a: "a = \<^bold>\<langle>un_Prim a\<^bold>\<rangle>"
using assms Nml_HcompD by metis
have b: "Nml b \<and> \<not> 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>\<Down> b)"
and "a \<^bold>\<Down> b \<in> VHom (a \<^bold>\<star> b) \<^bold>\<lfloor>a \<^bold>\<star> b\<^bold>\<rfloor>"
proof -
have 0: "\<And>b. \<lbrakk> Ide a \<and> Nml a; Ide b \<and> Nml b; Src a = Trg b \<rbrakk> \<Longrightarrow>
Can (a \<^bold>\<Down> b) \<and> a \<^bold>\<Down> b \<in> VHom (a \<^bold>\<star> b) \<^bold>\<lfloor>a \<^bold>\<star> b\<^bold>\<rfloor>"
proof (induct a, simp_all add: HcompNml_Nml_Src HcompNml_Trg_Nml)
fix x b
show "Ide b \<and> Nml b \<Longrightarrow> Can (Trg b \<^bold>\<Down> b) \<and> Arr (Trg b \<^bold>\<Down> b) \<and>
Dom (Trg b \<^bold>\<Down> b) = Trg b \<^bold>\<star> b \<and> Cod (Trg b \<^bold>\<Down> b) = b"
using Ide_implies_Can Ide_in_Hom Nmlize_Nml
apply (cases b, simp_all)
proof -
fix u v
assume uv: "Ide u \<and> Ide v \<and> Src u = Trg v \<and> Nml (u \<^bold>\<star> v)"
show "Can (Trg u \<^bold>\<Down> (u \<^bold>\<star> v)) \<and> Arr (Trg u \<^bold>\<Down> (u \<^bold>\<star> v)) \<and>
Dom (Trg u \<^bold>\<Down> (u \<^bold>\<star> v)) = Trg u \<^bold>\<star> u \<^bold>\<star> v \<and>
Cod (Trg u \<^bold>\<Down> (u \<^bold>\<star> v)) = u \<^bold>\<star> v"
using uv Ide_implies_Can Can_implies_Arr Ide_in_Hom
by (cases u, simp_all)
qed
show "ide x \<and> arr x \<Longrightarrow> Ide b \<and> Nml b \<Longrightarrow> \<^bold>\<langle>src x\<^bold>\<rangle>\<^sub>0 = Trg b \<Longrightarrow>
Can (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<Down> b) \<and> Arr (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<Down> b) \<and> Dom (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<Down> b) = \<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<star> b \<and> Cod (\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<Down> b) =
\<^bold>\<langle>x\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 \<and> Ide v \<and> Src u = Trg v \<and> Nml (u \<^bold>\<star> v)"
assume vw: "Src v = Trg w"
assume w: "Ide w \<and> Nml w"
assume I1: "\<And>w. \<lbrakk> Nml u; Ide w \<and> Nml w; Trg v = Trg w \<rbrakk> \<Longrightarrow>
Can (u \<^bold>\<Down> w) \<and> Arr (u \<^bold>\<Down> w) \<and>
Dom (u \<^bold>\<Down> w) = u \<^bold>\<star> w \<and> Cod (u \<^bold>\<Down> w) = u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w"
assume I2: "\<And>x. \<lbrakk> Nml v; Ide x \<and> Nml x; Trg w = Trg x \<rbrakk> \<Longrightarrow>
Can (v \<^bold>\<Down> x) \<and> Arr (v \<^bold>\<Down> x) \<and>
Dom (v \<^bold>\<Down> x) = v \<^bold>\<star> x \<and> Cod (v \<^bold>\<Down> x) = v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> x"
show "Can ((u \<^bold>\<star> v) \<^bold>\<Down> w) \<and> Arr ((u \<^bold>\<star> v) \<^bold>\<Down> w) \<and>
Dom ((u \<^bold>\<star> v) \<^bold>\<Down> w) = (u \<^bold>\<star> v) \<^bold>\<star> w \<and>
Cod ((u \<^bold>\<star> v) \<^bold>\<Down> w) = (\<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w"
proof -
have u: "Nml u \<and> Ide u"
using uv Nml_HcompD by blast
have v: "Nml v \<and> Ide v"
using uv Nml_HcompD by blast
have "is_Prim\<^sub>0 w \<Longrightarrow> ?thesis"
proof -
assume 1: "is_Prim\<^sub>0 w"
have 2: "(u \<^bold>\<star> v) \<^bold>\<Down> w = \<^bold>\<r>\<^bold>[u \<^bold>\<star> v\<^bold>]"
using 1 by (cases w, simp_all)
have 3: "Can (u \<^bold>\<Down> v) \<and> Arr (u \<^bold>\<Down> v) \<and> Dom (u \<^bold>\<Down> v) = u \<^bold>\<star> v \<and> Cod (u \<^bold>\<Down> v) = u \<^bold>\<star> v"
using u v uv 1 2 I1 Nmlize_Nml Nmlize.simps(3) by metis
hence 4: "VSeq (u \<^bold>\<Down> v) \<^bold>\<r>\<^bold>[u \<^bold>\<star> 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>\<star> v) \<^bold>\<Down> w)"
using 1 2 3 4 uv by (simp add: Ide_implies_Can)
moreover have "Dom ((u \<^bold>\<star> v) \<^bold>\<Down> w) = (u \<^bold>\<star> v) \<^bold>\<star> 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>\<star> v) \<^bold>\<Down> w) = \<^bold>\<lfloor>(u \<^bold>\<star> v) \<^bold>\<star> w\<^bold>\<rfloor>"
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 "\<not> is_Prim\<^sub>0 w \<Longrightarrow> ?thesis"
proof -
assume 1: "\<not> is_Prim\<^sub>0 w"
have 2: "(u \<^bold>\<star> v) \<^bold>\<Down> w = (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) \<^bold>\<cdot> (u \<^bold>\<star> v \<^bold>\<Down> w) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[u, v, w\<^bold>]"
using 1 u v uv w by (cases w; simp)
have 3: "Can (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) \<and> Dom (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) = u \<^bold>\<star> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor> \<and>
Cod (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) = \<^bold>\<lfloor>u \<^bold>\<star> (v \<^bold>\<star> w)\<^bold>\<rfloor>"
proof -
have "Can (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) \<and> Dom (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) = u \<^bold>\<star> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor> \<and>
Cod (u \<^bold>\<Down> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>) = u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>"
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>\<star> w"] Nml_Nmlize Ide_Nmlize_Ide
by simp
moreover have "u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor> = \<^bold>\<lfloor>u \<^bold>\<star> (v \<^bold>\<star> w)\<^bold>\<rfloor>"
using uv u w Nmlize_Hcomp Nmlize_Nmlize Nml_implies_Arr by simp
ultimately show ?thesis by presburger
qed
have 4: "Can (v \<^bold>\<Down> w) \<and> Dom (v \<^bold>\<Down> w) = v \<^bold>\<star> w \<and> Cod (v \<^bold>\<Down> w) = \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>"
using v w vw 1 2 I2 by simp
hence 5: "Src (v \<^bold>\<Down> w) = Src w \<and> Trg (v \<^bold>\<Down> w) = Trg v"
using Src_Dom Trg_Dom Can_implies_Arr by fastforce
have "Can (u \<^bold>\<star> (v \<^bold>\<Down> w)) \<and> Dom (u \<^bold>\<star> (v \<^bold>\<Down> w)) = u \<^bold>\<star> (v \<^bold>\<star> w) \<and>
Cod (u \<^bold>\<star> (v \<^bold>\<Down> w)) = u \<^bold>\<star> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>"
using u uv vw 4 5 Ide_implies_Can Ide_in_Hom by simp
moreover have "\<^bold>\<lfloor>u \<^bold>\<star> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>\<^bold>\<rfloor> = \<^bold>\<lfloor>u \<^bold>\<star> v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>u \<^bold>\<star> \<^bold>\<lfloor>v \<^bold>\<star> w\<^bold>\<rfloor>\<^bold>\<rfloor> = u \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> v) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w"
using u v w uv vw HcompNml_assoc by metis
also have "... = \<^bold>\<lfloor>u \<^bold>\<star> v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>"
using u v w by (metis Nmlize.simps(3) Nmlize_Nml)
finally show ?thesis by blast
qed
moreover have "Can \<^bold>\<a>\<^bold>[u, v, w\<^bold>] \<and> Dom \<^bold>\<a>\<^bold>[u, v, w\<^bold>] = (u \<^bold>\<star> v) \<^bold>\<star> w \<and>
Cod \<^bold>\<a>\<^bold>[u, v, w\<^bold>] = u \<^bold>\<star> (v \<^bold>\<star> 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>\<Down> b)" using assms 0 by blast
show "a \<^bold>\<Down> b \<in> VHom (a \<^bold>\<star> b) \<^bold>\<lfloor>a \<^bold>\<star> b\<^bold>\<rfloor>" 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>\<Down> v \<in> HHom (Src v) (Trg u)" and "u \<^bold>\<Down> v \<in> VHom (u \<^bold>\<star> v) \<^bold>\<lfloor>u \<^bold>\<star> v\<^bold>\<rfloor>"
proof -
show 1: "u \<^bold>\<Down> v \<in> VHom (u \<^bold>\<star> v) \<^bold>\<lfloor>u \<^bold>\<star> v\<^bold>\<rfloor>"
using assms Can_red2 Can_implies_Arr by simp
show "u \<^bold>\<Down> v \<in> HHom (Src v) (Trg u)"
using assms 1 Src_Dom [of "u \<^bold>\<Down> v"] Trg_Dom [of "u \<^bold>\<Down> 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>\<Down> v) = Src v" and "Trg (u \<^bold>\<Down> v) = Trg u"
and "Dom (u \<^bold>\<Down> v) = u \<^bold>\<star> v" and "Cod (u \<^bold>\<Down> v) = \<^bold>\<lfloor>u \<^bold>\<star> v\<^bold>\<rfloor>"
using assms red2_in_Hom by auto
lemma Can_red:
assumes "Ide u"
shows "Can (u\<^bold>\<down>)" and "u\<^bold>\<down> \<in> VHom u \<^bold>\<lfloor>u\<^bold>\<rfloor>"
proof -
have 0: "Ide u \<Longrightarrow> Can (u\<^bold>\<down>) \<and> u\<^bold>\<down> \<in> VHom u \<^bold>\<lfloor>u\<^bold>\<rfloor>"
proof (induct u, simp_all add: Dom_Ide Cod_Ide)
fix v w
assume v: "Can (v\<^bold>\<down>) \<and> Arr (v\<^bold>\<down>) \<and> Dom (v\<^bold>\<down>) = v \<and> Cod (v\<^bold>\<down>) = \<^bold>\<lfloor>v\<^bold>\<rfloor>"
assume w: "Can (w\<^bold>\<down>) \<and> Arr (w\<^bold>\<down>) \<and> Dom (w\<^bold>\<down>) = w \<and> Cod (w\<^bold>\<down>) = \<^bold>\<lfloor>w\<^bold>\<rfloor>"
assume vw: "Ide v \<and> Ide w \<and> Src v = Trg w"
show "(Nml (v \<^bold>\<star> w) \<longrightarrow>
Can v \<and> Can w \<and> v \<^bold>\<star> w = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and>
(\<not> Nml (v \<^bold>\<star> w) \<longrightarrow>
Can (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and> Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>) \<and>
Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and> Arr (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and> Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>) \<and>
Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and> Cod (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>)"
proof
show "Nml (v \<^bold>\<star> w) \<longrightarrow>
Can v \<and> Can w \<and> v \<^bold>\<star> w = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>"
using vw Nml_HcompD Ide_implies_Can Dom_Inv VcompNml_Ide_Nml Inv_Ide
Nmlize.simps(3) Nmlize_Nml
by metis
show "\<not> Nml (v \<^bold>\<star> w) \<longrightarrow>
Can (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and> Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>) \<and>
Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and> Arr (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and> Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>) \<and>
Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and> Cod (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>"
proof
assume 1: "\<not> Nml (v \<^bold>\<star> w)"
have "Can (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>)"
using v w vw Can_red2 Nml_Nmlize Ide_Nmlize_Ide Nml_HcompNml Ide_HcompNml
by simp
moreover have "Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>)"
using v w vw Src_Dom Trg_Dom by metis
moreover have "Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and> Cod (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>"
using v w vw Can_red2 Nml_Nmlize Ide_Nmlize_Ide by simp
ultimately show "Can (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and> Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>) \<and>
Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and> Arr (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) \<and>
Src (v\<^bold>\<down>) = Trg (w\<^bold>\<down>) \<and> Dom (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>w\<^bold>\<rfloor> \<and>
Cod (\<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>w\<^bold>\<rfloor>) = \<^bold>\<lfloor>v\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>w\<^bold>\<rfloor>"
using Can_implies_Arr by blast
qed
qed
qed
show "Can (u\<^bold>\<down>)" using assms 0 by blast
show "u\<^bold>\<down> \<in> VHom u \<^bold>\<lfloor>u\<^bold>\<rfloor>" using assms 0 by blast
qed
lemma red_in_Hom [intro]:
assumes "Ide t"
shows "t\<^bold>\<down> \<in> HHom (Src t) (Trg t)" and "t\<^bold>\<down> \<in> VHom t \<^bold>\<lfloor>t\<^bold>\<rfloor>"
proof -
show 1: "t\<^bold>\<down> \<in> VHom t \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Can_red Can_implies_Arr by simp
show "t\<^bold>\<down> \<in> HHom (Src t) (Trg t)"
using assms 1 Src_Dom [of "t\<^bold>\<down>"] Trg_Dom [of "t\<^bold>\<down>"] Can_red Can_implies_Arr by simp
qed
lemma red_simps [simp]:
assumes "Ide t"
shows "Src (t\<^bold>\<down>) = Src t" and "Trg (t\<^bold>\<down>) = Trg t"
and "Dom (t\<^bold>\<down>) = t" and "Cod (t\<^bold>\<down>) = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms red_in_Hom by auto
lemma red_Src:
assumes "Ide t"
shows "Src t\<^bold>\<down> = 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>\<down> = Trg t"
using assms is_Prim0_Trg [of t]
by (cases "Trg t", simp_all)
lemma Nmlize_red [simp]:
assumes "Ide t"
shows "\<^bold>\<lfloor>t\<^bold>\<down>\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
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>\<lfloor>t \<^bold>\<Down> u\<^bold>\<rfloor> = \<^bold>\<lfloor>t \<^bold>\<star> u\<^bold>\<rfloor>"
using assms Can_red2 Ide_Nmlize_Can Nmlize_in_Hom [of "t \<^bold>\<Down> u"] red2_in_Hom Ide_in_Hom
by simp
end
subsection "Evaluation"
text \<open>
The following locale is concerned with the evaluation of terms of the bicategorical
language determined by \<open>C\<close>, \<open>src\<^sub>C\<close>, and \<open>trg\<^sub>C\<close> in a bicategory \<open>(V, H, \<a>, \<i>, src, trg)\<close>,
given a source and target-preserving functor from \<open>C\<close> to \<open>V\<close>.
\<close>
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 \<a> \<i> src trg +
E: "functor" C V E
for C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V :: "'b comp" (infixr "\<cdot>" 55)
and H :: "'b comp" (infixr "\<star>" 53)
and \<a> :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>[_, _, _]")
and \<i> :: "'b \<Rightarrow> 'b" ("\<i>[_]")
and src :: "'b \<Rightarrow> 'b"
and trg :: "'b \<Rightarrow> 'b"
and E :: "'c \<Rightarrow> '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>\<lfloor>_\<^bold>\<rfloor>")
notation HcompNml (infixr "\<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor>" 53)
notation VcompNml (infixr "\<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor>" 55)
notation red ("_\<^bold>\<down>" [56] 56)
notation red2 (infixr "\<^bold>\<Down>" 53)
primrec eval :: "'c term \<Rightarrow> 'b" ("\<lbrace>_\<rbrace>")
where "\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle>\<^sub>0\<rbrace> = E f"
| "\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle>\<rbrace> = E f"
| "\<lbrace>t \<^bold>\<star> u\<rbrace> = \<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>"
| "\<lbrace>t \<^bold>\<cdot> u\<rbrace> = \<lbrace>t\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
| "\<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> = \<ll> \<lbrace>t\<rbrace>"
| "\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<ll>'.map \<lbrace>t\<rbrace>"
| "\<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> = \<rr> \<lbrace>t\<rbrace>"
| "\<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<rr>'.map \<lbrace>t\<rbrace>"
| "\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = \<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
| "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = \<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
lemma preserves_obj:
assumes "C.obj a"
shows "obj (E a)"
proof (unfold obj_def)
show "arr (E a) \<and> 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 \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
proof (induct t)
show "\<And>x. Arr \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace> : \<lbrace>Src \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace> : \<lbrace>Dom \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace>\<guillemotright>"
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 "\<And>x. Arr \<^bold>\<langle>x\<^bold>\<rangle> \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<rbrace> : \<lbrace>Src \<^bold>\<langle>x\<^bold>\<rangle>\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<langle>x\<^bold>\<rangle>\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<rbrace> : \<lbrace>Dom \<^bold>\<langle>x\<^bold>\<rangle>\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<langle>x\<^bold>\<rangle>\<rbrace>\<guillemotright>"
by (auto simp add: preserves_src preserves_trg)
show "\<And>t1 t2.
(Arr t1 \<Longrightarrow> \<guillemotleft>\<lbrace>t1\<rbrace> : \<lbrace>Src t1\<rbrace> \<rightarrow> \<lbrace>Trg t1\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t1\<rbrace> : \<lbrace>Dom t1\<rbrace> \<Rightarrow> \<lbrace>Cod t1\<rbrace>\<guillemotright>) \<Longrightarrow>
(Arr t2 \<Longrightarrow> \<guillemotleft>\<lbrace>t2\<rbrace> : \<lbrace>Src t2\<rbrace> \<rightarrow> \<lbrace>Trg t2\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t2\<rbrace> : \<lbrace>Dom t2\<rbrace> \<Rightarrow> \<lbrace>Cod t2\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr (t1 \<^bold>\<star> t2) \<Longrightarrow>
\<guillemotleft>\<lbrace>t1 \<^bold>\<star> t2\<rbrace> : \<lbrace>Src (t1 \<^bold>\<star> t2)\<rbrace> \<rightarrow> \<lbrace>Trg (t1 \<^bold>\<star> t2)\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<lbrace>t1 \<^bold>\<star> t2\<rbrace> : \<lbrace>Dom (t1 \<^bold>\<star> t2)\<rbrace> \<Rightarrow> \<lbrace>Cod (t1 \<^bold>\<star> t2)\<rbrace>\<guillemotright>"
using hcomp_in_hhom in_hhom_def vconn_implies_hpar(1) vconn_implies_hpar(2) by auto
show "\<And>t1 t2.
(Arr t1 \<Longrightarrow> \<guillemotleft>\<lbrace>t1\<rbrace> : \<lbrace>Src t1\<rbrace> \<rightarrow> \<lbrace>Trg t1\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t1\<rbrace> : \<lbrace>Dom t1\<rbrace> \<Rightarrow> \<lbrace>Cod t1\<rbrace>\<guillemotright>) \<Longrightarrow>
(Arr t2 \<Longrightarrow> \<guillemotleft>\<lbrace>t2\<rbrace> : \<lbrace>Src t2\<rbrace> \<rightarrow> \<lbrace>Trg t2\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t2\<rbrace> : \<lbrace>Dom t2\<rbrace> \<Rightarrow> \<lbrace>Cod t2\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr (t1 \<^bold>\<cdot> t2) \<Longrightarrow>
\<guillemotleft>\<lbrace>t1 \<^bold>\<cdot> t2\<rbrace> : \<lbrace>Src (t1 \<^bold>\<cdot> t2)\<rbrace> \<rightarrow> \<lbrace>Trg (t1 \<^bold>\<cdot> t2)\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<lbrace>t1 \<^bold>\<cdot> t2\<rbrace> : \<lbrace>Dom (t1 \<^bold>\<cdot> t2)\<rbrace> \<Rightarrow> \<lbrace>Cod (t1 \<^bold>\<cdot> t2)\<rbrace>\<guillemotright>"
using VSeq_implies_HPar seqI' by auto
show "\<And>t. (Arr t \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr \<^bold>\<l>\<^bold>[t\<^bold>] \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Src \<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Dom \<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace>\<guillemotright>"
proof (simp add: preserves_src preserves_trg)
fix t
assume t: "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
assume 1: "Arr t"
show "\<guillemotleft>\<ll> \<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<ll> \<lbrace>t\<rbrace> : \<lbrace>Trg t\<rbrace> \<star> \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
proof -
have "src (\<ll> \<lbrace>t\<rbrace>) = \<lbrace>Src t\<rbrace>"
using t 1
by (metis (no_types, lifting) \<ll>.preserves_cod arr_cod in_hhomE map_simp src_cod)
moreover have "trg (\<ll> \<lbrace>t\<rbrace>) = \<lbrace>Trg t\<rbrace>"
using t 1
by (metis (no_types, lifting) \<ll>.preserves_cod arr_cod in_hhomE map_simp trg_cod)
moreover have "\<guillemotleft>\<ll> \<lbrace>t\<rbrace> : \<lbrace>Trg t\<rbrace> \<star> \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
using t 1
apply (elim conjE in_hhomE)
by (intro in_homI, auto)
ultimately show ?thesis by auto
qed
qed
show "\<And>t. (Arr t \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Src \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Dom \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>\<guillemotright>"
proof (simp add: preserves_src preserves_trg)
fix t
assume t: "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
assume 1: "Arr t"
show "\<guillemotleft>\<ll>'.map \<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<ll>'.map \<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Trg t\<rbrace> \<star> \<lbrace>Cod t\<rbrace>\<guillemotright>"
proof -
have "src (\<ll>'.map \<lbrace>t\<rbrace>) = \<lbrace>Src t\<rbrace>"
using t 1 \<ll>'.preserves_dom arr_dom map_simp \<ll>'.preserves_reflects_arr src_dom
by (metis (no_types, lifting) in_hhomE)
moreover have "trg (\<ll>'.map \<lbrace>t\<rbrace>) = \<lbrace>Trg t\<rbrace>"
using t 1 \<ll>'.preserves_dom arr_dom map_simp \<ll>'.preserves_reflects_arr trg_dom
by (metis (no_types, lifting) in_hhomE)
moreover have "\<guillemotleft>\<ll>'.map \<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Trg t\<rbrace> \<star> \<lbrace>Cod t\<rbrace>\<guillemotright>"
using t 1 \<ll>'.preserves_hom
apply (intro in_homI)
apply auto[1]
apply fastforce
by fastforce
ultimately show ?thesis by blast
qed
qed
show "\<And>t. (Arr t \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr \<^bold>\<r>\<^bold>[t\<^bold>] \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Src \<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Dom \<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace>\<guillemotright>"
proof (simp add: preserves_src preserves_trg)
fix t
assume t: "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
assume 1: "Arr t"
show "\<guillemotleft>\<rr> \<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<rr> \<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<star> \<lbrace>Src t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
proof -
have "src (\<rr> \<lbrace>t\<rbrace>) = \<lbrace>Src t\<rbrace>"
using t 1 \<rr>.preserves_cod arr_cod map_simp \<rr>.preserves_reflects_arr src_cod
by (metis (no_types, lifting) in_hhomE)
moreover have "trg (\<rr> \<lbrace>t\<rbrace>) = \<lbrace>Trg t\<rbrace>"
using t 1 \<rr>.preserves_cod arr_cod map_simp \<rr>.preserves_reflects_arr trg_cod
by (metis (no_types, lifting) in_hhomE)
moreover have "\<guillemotleft>\<rr> \<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<star> \<lbrace>Src t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
using t 1 by force
ultimately show ?thesis by blast
qed
qed
show "\<And>t. (Arr t \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Src \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> : \<lbrace>Dom \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>\<guillemotright>"
proof (simp add: preserves_src preserves_trg)
fix t
assume t: "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
assume 1: "Arr t"
show "\<guillemotleft>\<rr>'.map \<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<rr>'.map \<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace> \<star> \<lbrace>Src t\<rbrace>\<guillemotright>"
proof -
have "src (\<rr>'.map \<lbrace>t\<rbrace>) = \<lbrace>Src t\<rbrace>"
using t 1 \<rr>'.preserves_dom arr_dom map_simp \<rr>'.preserves_reflects_arr src_dom
by (metis (no_types, lifting) in_hhomE)
moreover have "trg (\<rr>'.map \<lbrace>t\<rbrace>) = \<lbrace>Trg t\<rbrace>"
using t 1 \<rr>'.preserves_dom arr_dom map_simp \<rr>'.preserves_reflects_arr trg_dom
by (metis (no_types, lifting) in_hhomE)
moreover have "\<guillemotleft>\<rr>'.map \<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace> \<star> \<lbrace>Src t\<rbrace>\<guillemotright>"
using t 1 src_cod arr_cod \<rr>'.preserves_hom [of "\<lbrace>t\<rbrace>" "\<lbrace>Dom t\<rbrace>" "\<lbrace>Cod t\<rbrace>"]
apply (elim conjE in_hhomE)
by (intro in_homI, auto)
ultimately show ?thesis by blast
qed
qed
show "\<And>t u v.
(Arr t \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>) \<Longrightarrow>
(Arr u \<Longrightarrow> \<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Src u\<rbrace> \<rightarrow> \<lbrace>Trg u\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Dom u\<rbrace> \<Rightarrow> \<lbrace>Cod u\<rbrace>\<guillemotright>) \<Longrightarrow>
(Arr v \<Longrightarrow> \<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg v\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Dom v\<rbrace> \<Rightarrow> \<lbrace>Cod v\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr \<^bold>\<a>\<^bold>[t, u, v\<^bold>] \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> : \<lbrace>Src \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> : \<lbrace>Dom \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace>\<guillemotright>"
proof (simp add: preserves_src preserves_trg)
fix t u v
assume t: "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Trg u\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
assume u: "\<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Trg v\<rbrace> \<rightarrow> \<lbrace>Trg u\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Dom u\<rbrace> \<Rightarrow> \<lbrace>Cod u\<rbrace>\<guillemotright>"
assume v: "\<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg v\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Dom v\<rbrace> \<Rightarrow> \<lbrace>Cod v\<rbrace>\<guillemotright>"
assume tuv: "Arr t \<and> Arr u \<and> Arr v \<and> Src t = Trg u \<and> Src u = Trg v"
show "\<guillemotleft>\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) :
(\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace>) \<star> \<lbrace>Dom v\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace> \<star> \<lbrace>Cod u\<rbrace> \<star> \<lbrace>Cod v\<rbrace>\<guillemotright>"
proof -
have 1: "VVV.in_hom (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)
(\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>) (\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"
proof -
have "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) \<in>
VxVxV.hom (\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>) (\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"
using t u v tuv by simp
moreover have "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) \<in>
{\<tau>\<mu>\<nu>. arr (fst \<tau>\<mu>\<nu>) \<and> VV.arr (snd \<tau>\<mu>\<nu>) \<and>
src (fst \<tau>\<mu>\<nu>) = trg (fst (snd \<tau>\<mu>\<nu>))}"
using t u v tuv by fastforce
ultimately show ?thesis
using VVV.hom_char
[of "(\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>)" "(\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"]
by blast
qed
have 4: "VVV.arr (\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>)"
using 1 VVV.ide_dom VVV.dom_simp by (elim VVV.in_homE) force
have 5: "VVV.arr (\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"
- using 1 VVV.ide_cod VVV.cod_simp VVV.in_hom_char by blast
+ using 1 VVV.ide_cod VVV.cod_simp VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C by blast
have 2: "\<guillemotleft>\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) :
(\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace>) \<star> \<lbrace>Dom v\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace> \<star> \<lbrace>Cod u\<rbrace> \<star> \<lbrace>Cod v\<rbrace>\<guillemotright>"
using 1 4 5 HoHV_def HoVH_def \<alpha>_def
\<alpha>.preserves_hom [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)" "(\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>)"
"(\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"]
by simp
have 3: "\<guillemotleft>\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright>"
proof
show "arr (\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>))"
using 2 by auto
show "src (\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = \<lbrace>Src v\<rbrace>"
proof -
have "src (\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = src ((\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace>) \<star> \<lbrace>Dom v\<rbrace>)"
using 2 src_dom [of "\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] by fastforce
also have "... = src \<lbrace>Dom v\<rbrace>"
- using 4 VVV.arr_char VV.arr_char by simp
+ using 4 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = src (dom \<lbrace>v\<rbrace>)"
using v by fastforce
also have "... = \<lbrace>Src v\<rbrace>"
using v by auto
finally show ?thesis by auto
qed
show "trg (\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = \<lbrace>Trg t\<rbrace>"
proof -
have "trg (\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = trg ((\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace>) \<star> \<lbrace>Dom v\<rbrace>)"
using 2 trg_dom [of "\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] by fastforce
also have "... = trg \<lbrace>Dom t\<rbrace>"
- using 4 VVV.arr_char VV.arr_char by simp
+ using 4 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = trg (dom \<lbrace>t\<rbrace>)"
using t by fastforce
also have "... = \<lbrace>Trg t\<rbrace>"
using t by auto
finally show ?thesis by auto
qed
qed
show ?thesis using 2 3 by simp
qed
qed
show "\<And>t u v.
(Arr t \<Longrightarrow> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>) \<Longrightarrow>
(Arr u \<Longrightarrow> \<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Src u\<rbrace> \<rightarrow> \<lbrace>Trg u\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Dom u\<rbrace> \<Rightarrow> \<lbrace>Cod u\<rbrace>\<guillemotright>) \<Longrightarrow>
(Arr v \<Longrightarrow> \<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg v\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Dom v\<rbrace> \<Rightarrow> \<lbrace>Cod v\<rbrace>\<guillemotright>) \<Longrightarrow>
Arr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \<Longrightarrow>
\<guillemotleft>\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> : \<lbrace>Src \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> \<rightarrow> \<lbrace>Trg \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> : \<lbrace>Dom \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> \<Rightarrow> \<lbrace>Cod \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace>\<guillemotright>"
proof (simp add: preserves_src preserves_trg)
fix t u v
assume t: "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Trg u\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
assume u: "\<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Trg v\<rbrace> \<rightarrow> \<lbrace>Trg u\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>u\<rbrace> : \<lbrace>Dom u\<rbrace> \<Rightarrow> \<lbrace>Cod u\<rbrace>\<guillemotright>"
assume v: "\<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg v\<rbrace>\<guillemotright> \<and> \<guillemotleft>\<lbrace>v\<rbrace> : \<lbrace>Dom v\<rbrace> \<Rightarrow> \<lbrace>Cod v\<rbrace>\<guillemotright>"
assume tuv: "Arr t \<and> Arr u \<and> Arr v \<and> Src t = Trg u \<and> Src u = Trg v"
show "\<guillemotleft>\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright> \<and>
\<guillemotleft>\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) :
\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace> \<star> \<lbrace>Dom v\<rbrace> \<Rightarrow> (\<lbrace>Cod t\<rbrace> \<star> \<lbrace>Cod u\<rbrace>) \<star> \<lbrace>Cod v\<rbrace>\<guillemotright>"
proof -
have 1: "VVV.in_hom (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)
(\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>) (\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"
- using t u v tuv VVV.hom_char VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char
+ using t u v tuv VVV.hom_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
apply (elim conjE in_hhomE in_homE)
apply (intro VVV.in_homI)
by simp_all
have 4: "VVV.arr (\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>)"
- using "1" VVV.in_hom_char by blast
+ using "1" VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C by blast
have 5: "VVV.arr (\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"
- using "1" VVV.in_hom_char by blast
+ using "1" VVV.in_hom_char\<^sub>S\<^sub>b\<^sub>C by blast
have 2: "\<guillemotleft>\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) :
\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace> \<star> \<lbrace>Dom v\<rbrace> \<Rightarrow> (\<lbrace>Cod t\<rbrace> \<star> \<lbrace>Cod u\<rbrace>) \<star> \<lbrace>Cod v\<rbrace>\<guillemotright>"
using 1 4 5 HoHV_def HoVH_def \<alpha>'.map_def
\<alpha>'.preserves_hom [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)" "(\<lbrace>Dom t\<rbrace>, \<lbrace>Dom u\<rbrace>, \<lbrace>Dom v\<rbrace>)"
"(\<lbrace>Cod t\<rbrace>, \<lbrace>Cod u\<rbrace>, \<lbrace>Cod v\<rbrace>)"]
by simp
have 3: "\<guillemotleft>\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>) : \<lbrace>Src v\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright>"
proof
show "arr (\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>))"
using 2 by auto
show "src (\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = \<lbrace>Src v\<rbrace>"
proof -
have "src (\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = src (\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace> \<star> \<lbrace>Dom v\<rbrace>)"
using 2 src_dom [of "\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] by fastforce
also have "... = src \<lbrace>Dom v\<rbrace>"
- using 4 VVV.arr_char VV.arr_char by simp
+ using 4 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = src (dom \<lbrace>v\<rbrace>)"
using v by fastforce
also have "... = \<lbrace>Src v\<rbrace>"
using v by auto
finally show ?thesis by auto
qed
show "trg (\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = \<lbrace>Trg t\<rbrace>"
proof -
have "trg (\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) = trg (\<lbrace>Dom t\<rbrace> \<star> \<lbrace>Dom u\<rbrace> \<star> \<lbrace>Dom v\<rbrace>)"
using 2 trg_dom [of "\<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] by fastforce
also have "... = trg \<lbrace>Dom t\<rbrace>"
- using 4 VVV.arr_char VV.arr_char hseqI' by simp
+ using 4 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseqI' by simp
also have "... = trg (dom \<lbrace>t\<rbrace>)"
using t by fastforce
also have "... = \<lbrace>Trg t\<rbrace>"
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 "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Src t\<rbrace> \<rightarrow> \<lbrace>Trg t\<rbrace>\<guillemotright>" and "\<guillemotleft>\<lbrace>t\<rbrace> : \<lbrace>Dom t\<rbrace> \<Rightarrow> \<lbrace>Cod t\<rbrace>\<guillemotright>"
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 \<lbrace>f\<rbrace>" and "\<lbrace>Src f\<rbrace> = src \<lbrace>f\<rbrace>" and "\<lbrace>Trg f\<rbrace> = trg \<lbrace>f\<rbrace>"
and "\<lbrace>Dom f\<rbrace> = dom \<lbrace>f\<rbrace>" and "\<lbrace>Cod f\<rbrace> = cod \<lbrace>f\<rbrace>"
using assms eval_in_hom [of f] by auto
lemma eval_simps':
assumes "Arr f"
shows "arr \<lbrace>f\<rbrace>" and "src \<lbrace>f\<rbrace> = \<lbrace>Src f\<rbrace>" and "trg \<lbrace>f\<rbrace> = \<lbrace>Trg f\<rbrace>"
and "dom \<lbrace>f\<rbrace> = \<lbrace>Dom f\<rbrace>" and "cod \<lbrace>f\<rbrace> = \<lbrace>Cod f\<rbrace>"
using assms eval_in_hom by auto
lemma obj_eval_Obj:
shows "Obj t \<Longrightarrow> obj \<lbrace>t\<rbrace>"
using preserves_obj
by (induct t) auto
lemma ide_eval_Ide:
shows "Ide t \<Longrightarrow> ide \<lbrace>t\<rbrace>"
by (induct t, auto simp add: eval_simps')
lemma arr_eval_Arr [simp]:
assumes "Arr t"
shows "arr \<lbrace>t\<rbrace>"
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 "\<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> = \<l>[\<lbrace>Cod t\<rbrace>] \<cdot> (trg \<lbrace>t\<rbrace> \<star> \<lbrace>t\<rbrace>)"
using assms \<ll>.is_natural_2 \<ll>_ide_simp by (simp add: eval_simps)
lemma eval_Lunit' [simp]:
assumes "Arr t"
shows "\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<l>\<^sup>-\<^sup>1[\<lbrace>Cod t\<rbrace>] \<cdot> \<lbrace>t\<rbrace>"
using assms \<ll>'.is_natural_2 \<ll>_ide_simp by (simp add: eval_simps)
lemma eval_Runit [simp]:
assumes "Arr t"
shows "\<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> = \<r>[\<lbrace>Cod t\<rbrace>] \<cdot> (\<lbrace>t\<rbrace> \<star> src \<lbrace>t\<rbrace>)"
using assms \<rr>.is_natural_2 \<rr>_ide_simp by (simp add: eval_simps)
lemma eval_Runit' [simp]:
assumes "Arr t"
shows "\<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<r>\<^sup>-\<^sup>1[\<lbrace>Cod t\<rbrace>] \<cdot> \<lbrace>t\<rbrace>"
using assms \<rr>'.is_natural_2 \<rr>_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 "\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = \<alpha> (cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>) \<cdot> ((\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>) \<star> \<lbrace>v\<rbrace>)"
proof -
have "\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = \<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)" by simp
also have "... = \<alpha> (VVV.cod (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) \<cdot> HoHV (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
- using assms \<alpha>.is_natural_2 [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] VVV.arr_char VVV.cod_char
+ using assms \<alpha>.is_natural_2 [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
\<alpha>.is_extensional \<alpha>_def
by auto
also have "... = \<alpha> (cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>) \<cdot> ((\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>) \<star> \<lbrace>v\<rbrace>)"
unfolding HoHV_def \<alpha>_def
- using assms VVV.arr_char VV.arr_char VVV.cod_char \<alpha>.is_extensional
+ using assms VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C \<alpha>.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 "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = \<a>\<^sup>-\<^sup>1[cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>] \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace> \<star> \<lbrace>v\<rbrace>)"
proof -
have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = \<alpha>'.map (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)" by simp
also have "... = \<alpha>'.map (VVV.cod (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)) \<cdot> HoVH (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
- using assms \<alpha>'.is_natural_2 [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] VVV.arr_char VVV.cod_char
+ using assms \<alpha>'.is_natural_2 [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
\<alpha>'.is_extensional
by simp
also have "... = \<alpha>'.map (cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>) \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace> \<star> \<lbrace>v\<rbrace>)"
unfolding HoVH_def
- using assms VVV.arr_char VV.arr_char VVV.cod_char \<alpha>'.is_extensional
+ using assms VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C \<alpha>'.is_extensional
apply simp
using eval_simps'(2) eval_simps'(3) by presburger
finally show ?thesis
using \<a>'_def by simp
qed
lemma eval_Lunit_Ide [simp]:
assumes "Ide t"
shows "\<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> = \<l>[\<lbrace>t\<rbrace>]"
using assms \<ll>_ide_simp ide_eval_Ide by simp
lemma eval_Lunit'_Ide [simp]:
assumes "Ide t"
shows "\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<l>\<^sup>-\<^sup>1[\<lbrace>t\<rbrace>]"
using assms \<ll>_ide_simp ide_eval_Ide by simp
lemma eval_Runit_Ide [simp]:
assumes "Ide t"
shows "\<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> = \<r>[\<lbrace>t\<rbrace>]"
using assms \<rr>_ide_simp ide_eval_Ide by simp
lemma eval_Runit'_Ide [simp]:
assumes "Ide t"
shows "\<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<r>\<^sup>-\<^sup>1[\<lbrace>t\<rbrace>]"
using assms \<rr>_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 "\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = \<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
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 "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = \<a>\<^sup>-\<^sup>1[\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>]"
using assms \<a>'_def by simp
lemma iso_eval_Can:
shows "Can t \<Longrightarrow> iso \<lbrace>t\<rbrace>"
proof (induct t; simp add: fsts.intros snds.intros)
show "\<And>x. C.obj x \<Longrightarrow> iso (E x)" by auto
show "\<And>t1 t2. \<lbrakk> iso \<lbrace>t1\<rbrace>; iso \<lbrace>t2\<rbrace>; Can t1 \<and> Can t2 \<and> Src t1 = Trg t2 \<rbrakk> \<Longrightarrow>
iso (\<lbrace>t1\<rbrace> \<star> \<lbrace>t2\<rbrace>)"
using Can_implies_Arr by (simp add: eval_simps')
show "\<And>t1 t2. \<lbrakk> iso \<lbrace>t1\<rbrace>; iso \<lbrace>t2\<rbrace>; Can t1 \<and> Can t2 \<and> Dom t1 = Cod t2 \<rbrakk> \<Longrightarrow>
iso (\<lbrace>t1\<rbrace> \<cdot> \<lbrace>t2\<rbrace>)"
using Can_implies_Arr isos_compose by (simp add: eval_simps')
show "\<And>t. \<lbrakk> iso \<lbrace>t\<rbrace>; Can t \<rbrakk> \<Longrightarrow> iso (\<ll> \<lbrace>t\<rbrace>)"
using \<ll>.preserves_iso by auto
show "\<And>t. \<lbrakk> iso \<lbrace>t\<rbrace>; Can t \<rbrakk> \<Longrightarrow> iso (\<ll>'.map \<lbrace>t\<rbrace>)"
using \<ll>'.preserves_iso by simp
show "\<And>t. \<lbrakk> iso \<lbrace>t\<rbrace>; Can t \<rbrakk> \<Longrightarrow> iso (\<rr> \<lbrace>t\<rbrace>)"
using \<rr>.preserves_iso by auto
show "\<And>t. \<lbrakk> iso \<lbrace>t\<rbrace>; Can t \<rbrakk> \<Longrightarrow> iso (\<rr>'.map \<lbrace>t\<rbrace>)"
using \<rr>'.preserves_iso by simp
fix t1 t2 t3
assume t1: "iso \<lbrace>t1\<rbrace>" and t2: "iso \<lbrace>t2\<rbrace>" and t3: "iso \<lbrace>t3\<rbrace>"
assume 1: "Can t1 \<and> Can t2 \<and> Can t3 \<and> Src t1 = Trg t2 \<and> Src t2 = Trg t3"
have 2: "VVV.iso (\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>)"
proof -
have 3: "VxVxV.iso (\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>)"
using t1 t2 t3 Can_implies_Arr VxVxV.iso_char VxV.iso_char by simp
moreover have "VVV.arr (VxVxV.inv (\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>))"
proof -
have "VxVxV.inv (\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>) = (inv \<lbrace>t1\<rbrace>, inv \<lbrace>t2\<rbrace>, inv \<lbrace>t3\<rbrace>)"
using t1 t2 t3 3 by simp
thus ?thesis
- using t1 t2 t3 1 Can_implies_Arr VVV.arr_char VV.arr_char
+ using t1 t2 t3 1 Can_implies_Arr VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (simp add: eval_simps')
qed
ultimately show ?thesis
- using t1 t2 t3 1 Can_implies_Arr VVV.iso_char VVV.arr_char
+ using t1 t2 t3 1 Can_implies_Arr VVV.iso_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (auto simp add: eval_simps')
qed
show "iso (\<alpha> (\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>))"
using 2 \<alpha>_def \<alpha>.preserves_iso by auto
show "iso (\<alpha>'.map (\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>))"
using 2 \<alpha>'.preserves_iso by simp
qed
lemma eval_Inv_Can:
shows "Can t \<Longrightarrow> \<lbrace>Inv t\<rbrace> = inv \<lbrace>t\<rbrace>"
proof (induct t)
show "\<And>x. Can \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 \<Longrightarrow> \<lbrace>Inv \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace> = inv \<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0\<rbrace>" by auto
show "\<And>x. Can \<^bold>\<langle>x\<^bold>\<rangle> \<Longrightarrow> \<lbrace>Inv \<^bold>\<langle>x\<^bold>\<rangle>\<rbrace> = inv \<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<rbrace>" by simp
show "\<And>t1 t2. (Can t1 \<Longrightarrow> \<lbrace>Inv t1\<rbrace> = inv \<lbrace>t1\<rbrace>) \<Longrightarrow>
(Can t2 \<Longrightarrow> \<lbrace>Inv t2\<rbrace> = inv \<lbrace>t2\<rbrace>) \<Longrightarrow>
Can (t1 \<^bold>\<star> t2) \<Longrightarrow> \<lbrace>Inv (t1 \<^bold>\<star> t2)\<rbrace> = inv \<lbrace>t1 \<^bold>\<star> t2\<rbrace>"
using iso_eval_Can Can_implies_Arr
by (simp add: eval_simps')
show "\<And>t1 t2. (Can t1 \<Longrightarrow> \<lbrace>Inv t1\<rbrace> = inv \<lbrace>t1\<rbrace>) \<Longrightarrow>
(Can t2 \<Longrightarrow> \<lbrace>Inv t2\<rbrace> = inv \<lbrace>t2\<rbrace>) \<Longrightarrow>
Can (t1 \<^bold>\<cdot> t2) \<Longrightarrow> \<lbrace>Inv (t1 \<^bold>\<cdot> t2)\<rbrace> = inv \<lbrace>t1 \<^bold>\<cdot> t2\<rbrace>"
using iso_eval_Can inv_comp Can_implies_Arr
by (simp add: eval_simps')
fix t
assume I: "Can t \<Longrightarrow> \<lbrace>Inv t\<rbrace> = inv \<lbrace>t\<rbrace>"
show "Can \<^bold>\<l>\<^bold>[t\<^bold>] \<Longrightarrow> \<lbrace>Inv \<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> = inv \<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace>"
proof -
assume t: "Can \<^bold>\<l>\<^bold>[t\<^bold>]"
have "\<lbrace>Inv \<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]\<rbrace>" by simp
also have "... = \<ll>'.map (inv \<lbrace>t\<rbrace>)"
using t I by simp
also have "... = \<ll>'.map (cod (inv \<lbrace>t\<rbrace>)) \<cdot> inv \<lbrace>t\<rbrace>"
using t \<ll>'.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 (\<lbrace>t\<rbrace> \<cdot> \<ll> (dom \<lbrace>t\<rbrace>))"
proof -
have 1: "iso \<lbrace>t\<rbrace>" using t iso_eval_Can by simp
moreover have "iso (\<ll> (dom \<lbrace>t\<rbrace>))"
using t 1 \<ll>.components_are_iso ide_dom by blast
moreover have "seq \<lbrace>t\<rbrace> (\<ll> (dom \<lbrace>t\<rbrace>))"
using t 1 iso_is_arr by auto
ultimately show ?thesis
using t 1 inv_comp by auto
qed
also have "... = inv \<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace>"
using t iso_eval_Can \<ll>_ide_simp lunit_naturality Can_implies_Arr eval_Lunit
by (auto simp add: eval_simps)
finally show ?thesis by blast
qed
show "Can \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow> \<lbrace>Inv \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = inv \<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>"
proof -
assume t: "Can (Lunit' t)"
have "\<lbrace>Inv (Lunit' t)\<rbrace> = \<lbrace>Lunit (Inv t)\<rbrace>" by simp
also have "... = \<ll> (inv \<lbrace>t\<rbrace>)"
using t I by simp
also have "... = inv \<lbrace>t\<rbrace> \<cdot> \<ll> (dom (inv \<lbrace>t\<rbrace>))"
using t \<ll>.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 (\<ll>'.map (cod \<lbrace>t\<rbrace>) \<cdot> \<lbrace>t\<rbrace>)"
proof -
have 1: "iso \<lbrace>t\<rbrace>" using t iso_eval_Can by simp
moreover have "iso (\<ll>'.map (cod \<lbrace>t\<rbrace>))"
using t 1 \<ll>'.components_are_iso ide_cod by blast
moreover have "seq (\<ll>'.map (cod \<lbrace>t\<rbrace>)) \<lbrace>t\<rbrace>"
using t 1 iso_is_arr by auto
ultimately show ?thesis
using t 1 inv_comp by auto
qed
also have "... = inv \<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>"
using t \<ll>'.is_natural_2 iso_eval_Can iso_is_arr by force
finally show ?thesis by auto
qed
show "Can \<^bold>\<r>\<^bold>[t\<^bold>] \<Longrightarrow> \<lbrace>Inv \<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> = inv \<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace>"
proof -
assume t: "Can \<^bold>\<r>\<^bold>[t\<^bold>]"
have "\<lbrace>Inv \<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]\<rbrace>" by simp
also have "... = \<rr>'.map (inv \<lbrace>t\<rbrace>)"
using t I by simp
also have "... = \<rr>'.map (cod (inv \<lbrace>t\<rbrace>)) \<cdot> inv \<lbrace>t\<rbrace>"
using t \<rr>'.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 (\<lbrace>t\<rbrace> \<cdot> \<rr> (dom \<lbrace>t\<rbrace>))"
proof -
have 1: "iso \<lbrace>t\<rbrace>" using t iso_eval_Can by simp
moreover have "iso (\<rr> (dom \<lbrace>t\<rbrace>))"
using t 1 \<rr>.components_are_iso ide_dom by blast
moreover have "seq \<lbrace>t\<rbrace> (\<rr> (dom \<lbrace>t\<rbrace>))"
using t 1 iso_is_arr by simp
ultimately show ?thesis
using t 1 inv_comp by auto
qed
also have "... = inv \<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace>"
using t \<rr>_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>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow> \<lbrace>Inv \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = inv \<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>"
proof -
assume t: "Can \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]"
have "\<lbrace>Inv \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<r>\<^bold>[Inv t\<^bold>]\<rbrace>"
by simp
also have "... = \<rr> (inv \<lbrace>t\<rbrace>)"
using t I by simp
also have "... = inv \<lbrace>t\<rbrace> \<cdot> \<rr> (dom (inv \<lbrace>t\<rbrace>))"
using t \<rr>.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 (\<rr>'.map (cod \<lbrace>t\<rbrace>) \<cdot> \<lbrace>t\<rbrace>)"
proof -
have 1: "iso \<lbrace>t\<rbrace>" using t iso_eval_Can by simp
moreover have "iso (\<rr>'.map (cod \<lbrace>t\<rbrace>))"
using t 1 \<rr>'.components_are_iso ide_cod by blast
moreover have "seq (\<rr>'.map (cod \<lbrace>t\<rbrace>)) \<lbrace>t\<rbrace>"
using t 1 iso_is_arr by auto
ultimately show ?thesis
using t 1 inv_comp by auto
qed
also have "... = inv \<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>"
using t \<rr>'.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 \<Longrightarrow> \<lbrace>Inv t\<rbrace> = inv \<lbrace>t\<rbrace>"
assume I2: "Can u \<Longrightarrow> \<lbrace>Inv u\<rbrace> = inv \<lbrace>u\<rbrace>"
assume I3: "Can v \<Longrightarrow> \<lbrace>Inv v\<rbrace> = inv \<lbrace>v\<rbrace>"
show "Can \<^bold>\<a>\<^bold>[t, u, v\<^bold>] \<Longrightarrow> \<lbrace>Inv \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = inv \<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace>"
proof -
assume "Can \<^bold>\<a>\<^bold>[t, u, v\<^bold>]"
hence tuv: "Can t \<and> Can u \<and> Can v \<and> Src t = Trg u \<and> Src u = Trg v" by simp
have "\<lbrace>Inv \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Inv t, Inv u, Inv v\<^bold>]\<rbrace>" by simp
also have "... = \<a>\<^sup>-\<^sup>1[dom \<lbrace>t\<rbrace>, dom \<lbrace>u\<rbrace>, dom \<lbrace>v\<rbrace>] \<cdot> (inv \<lbrace>t\<rbrace> \<star> inv \<lbrace>u\<rbrace> \<star> inv \<lbrace>v\<rbrace>)"
using tuv I1 I2 I3 eval_in_hom \<alpha>'.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 ((\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace> \<star> \<lbrace>v\<rbrace>) \<cdot> \<alpha> (dom \<lbrace>t\<rbrace>, dom \<lbrace>u\<rbrace>, dom \<lbrace>v\<rbrace>))"
using tuv iso_eval_Can Can_implies_Arr eval_simps'(2) eval_simps'(3) \<alpha>_def
by (simp add: inv_comp)
also have "... = inv (\<alpha> (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>))"
using tuv Can_implies_Arr \<alpha>_def
by (metis assoc_is_natural_1 arr_eval_Arr eval_simps'(2) eval_simps'(3) fst_conv snd_conv)
also have "... = inv \<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace>" by simp
finally show ?thesis by blast
qed
show "Can \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \<Longrightarrow> \<lbrace>Inv \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = inv \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace>"
proof -
assume tuv: "Can \<^bold>\<a>\<^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 "\<lbrace>Inv \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<a>\<^bold>[Inv t, Inv u, Inv v\<^bold>]\<rbrace>" by simp
also have "... = (inv \<lbrace>t\<rbrace> \<star> inv \<lbrace>u\<rbrace> \<star> inv \<lbrace>v\<rbrace>) \<cdot> \<alpha> (cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>)"
using \<alpha>_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 (\<a>\<^sup>-\<^sup>1[cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>] \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace> \<star> \<lbrace>v\<rbrace>))"
using tuv inv_comp [of "\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace> \<star> \<lbrace>v\<rbrace>" "\<a>\<^sup>-\<^sup>1[cod \<lbrace>t\<rbrace>, cod \<lbrace>u\<rbrace>, cod \<lbrace>v\<rbrace>]"]
Can_implies_Arr iso_assoc \<alpha>_def
by (simp add: eval_simps'(2) eval_simps'(3) iso_eval_Can)
also have 1: "... = inv (((\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>) \<star> \<lbrace>v\<rbrace>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<lbrace>t\<rbrace>, dom \<lbrace>u\<rbrace>, dom \<lbrace>v\<rbrace>])"
using tuv assoc'_naturality [of "\<lbrace>t\<rbrace>" "\<lbrace>u\<rbrace>" "\<lbrace>v\<rbrace>"] Can_implies_Arr
eval_simps'(2) eval_simps'(3)
by simp
also have "... = inv \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace>"
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 "\<lbrace>t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u\<rbrace> = \<lbrace>t\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
proof -
have "\<And>u. \<lbrakk> Nml t; Nml u; VSeq t u \<rbrakk> \<Longrightarrow> \<lbrace>t \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u\<rbrace> = \<lbrace>t\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
proof (induct t, simp_all add: eval_simps)
fix u a
assume u: "Nml u"
assume 1: "Arr u \<and> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 = Cod u"
show "\<lbrace>u\<rbrace> = cod \<lbrace>u\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
using 1 comp_cod_arr by simp
next
fix u f
assume u: "Nml u"
assume f: "C.arr f"
assume 1: "Arr u \<and> \<^bold>\<langle>C.dom f\<^bold>\<rangle> = Cod u"
show "\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u\<rbrace> = E f \<cdot> \<lbrace>u\<rbrace>"
using f u 1 as_nat_trans.preserves_comp_2 by (cases u; simp)
next
fix u v w
assume I1: "\<And>u. \<lbrakk> Nml v; Nml u; Arr u \<and> Dom v = Cod u \<rbrakk> \<Longrightarrow> \<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u\<rbrace> = \<lbrace>v\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
assume I2: "\<And>u. \<lbrakk> Nml w; Nml u; Arr u \<and> Dom w = Cod u \<rbrakk> \<Longrightarrow> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u\<rbrace> = \<lbrace>w\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
assume vw: "Nml (v \<^bold>\<star> w)"
have v: "Nml v \<and> 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 \<and> Arr w \<and> Src v = Trg w \<and> Arr u \<and> Dom v \<^bold>\<star> Dom w = Cod u"
show "\<lbrace>(v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> u\<rbrace> = (\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<cdot> \<lbrace>u\<rbrace>"
using u 1 HcompNml_in_Hom apply (cases u, simp_all)
proof -
fix x y
assume 3: "u = x \<^bold>\<star> y"
have x: "Nml x"
using u 1 3 Nml_HcompD by blast
have y: "Nml y"
using u x 1 3 Nml_HcompD by blast
assume 4: "Arr v \<and> Arr w \<and> Src v = Trg w \<and> Dom v = Cod x \<and> Dom w = Cod y"
have "\<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y\<rbrace> = \<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x \<^bold>\<star> w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y\<rbrace>"
using v w x y 4 HcompNml_in_Hom by simp
moreover have "... = \<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y\<rbrace>" by simp
moreover have "... = \<lbrace>v\<rbrace> \<cdot> \<lbrace>x\<rbrace> \<star> \<lbrace>w\<rbrace> \<cdot> \<lbrace>y\<rbrace>"
using v w x y 4 I1 [of x] I2 [of y] Nml_implies_Arr by simp
moreover have "... = (\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<cdot> (\<lbrace>x\<rbrace> \<star> \<lbrace>y\<rbrace>)"
using v w x y 4 Nml_implies_Arr interchange [of "\<lbrace>v\<rbrace>" "\<lbrace>x\<rbrace>" "\<lbrace>w\<rbrace>" "\<lbrace>y\<rbrace>"]
by (simp add: eval_simps')
ultimately have "\<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y\<rbrace> = (\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<cdot> (\<lbrace>x\<rbrace> \<star> \<lbrace>y\<rbrace>)" by presburger
moreover have "arr \<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x\<rbrace> \<and> arr \<lbrace>w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y\<rbrace>"
using v w x y 4 VcompNml_in_Hom by simp
ultimately show "\<lbrace>v \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> x\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> y\<rbrace> = (\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<cdot> (\<lbrace>x\<rbrace> \<star> \<lbrace>y\<rbrace>)"
by simp
qed
qed
thus ?thesis using assms by blast
qed
lemma eval_red_Hcomp:
assumes "Ide a" and "Ide b"
shows "\<lbrace>(a \<^bold>\<star> b)\<^bold>\<down>\<rbrace> = \<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>)"
proof -
have "Nml (a \<^bold>\<star> b) \<Longrightarrow> ?thesis"
proof -
assume 1: "Nml (a \<^bold>\<star> b)"
hence 2: "Nml a \<and> Nml b \<and> Src a = Trg b"
using Nml_HcompD(3-4,7) by simp
have "\<lbrace>(a \<^bold>\<star> b)\<^bold>\<down>\<rbrace> = \<lbrace>a\<rbrace> \<star> \<lbrace>b\<rbrace>"
using 1 Nml_HcompD
by (metis eval.simps(3) red_Nml)
also have "... = \<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>)"
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 "\<not> Nml (a \<^bold>\<star> b) \<Longrightarrow> ?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>\<langle>a\<^bold>\<rangle>\<^sub>0 is replaced by Src t? *)
lemma eval_red2_Nml_Prim\<^sub>0:
assumes "Ide t" and "Nml t" and "Src t = \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0"
shows "\<lbrace>t \<^bold>\<Down> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0\<rbrace> = \<r>[\<lbrace>t\<rbrace>]"
using assms \<rr>_ide_simp
apply (cases t)
apply simp_all
proof -
show "C.obj a \<Longrightarrow> t = \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<Longrightarrow> \<ll> (E a) = \<r>[E a]"
using unitor_coincidence obj_eval_Obj [of t] \<ll>_ide_simp by auto
show "\<And>b c. Ide b \<and> Ide c \<and> Src b = Trg c \<Longrightarrow> \<rr> (\<lbrace>b\<rbrace> \<star> \<lbrace>c\<rbrace>) = \<r>[\<lbrace>b\<rbrace> \<star> \<lbrace>c\<rbrace>]"
using \<rr>_ide_simp by (simp add: eval_simps' ide_eval_Ide)
qed
end
text \<open>
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.
\<close>
locale self_evaluation_map =
bicategory
begin
sublocale bicategorical_language V src trg ..
sublocale evaluation_map V src trg V H \<a> \<i> src trg \<open>\<lambda>\<mu>. if arr \<mu> then \<mu> else null\<close>
using src.is_extensional trg.is_extensional
by (unfold_locales, auto)
notation eval ("\<lbrace>_\<rbrace>")
notation Nmlize ("\<^bold>\<lfloor>_\<^bold>\<rfloor>")
end
subsection "Coherence"
text \<open>
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''.
\<close>
context evaluation_map
begin
abbreviation coherent
where "coherent t \<equiv> \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace> = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>"
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 \<longleftrightarrow> \<lbrace>t\<rbrace> = inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>"
proof
assume 1: "coherent t"
have "inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace> = inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace>"
using 1 by simp
also have "... = (inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>Cod t\<^bold>\<down>\<rbrace>) \<cdot> \<lbrace>t\<rbrace>"
using comp_assoc by simp
also have "... = \<lbrace>t\<rbrace>"
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 "\<lbrace>t\<rbrace> = inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>"
by presburger
next
assume 1: "\<lbrace>t\<rbrace> = inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>"
hence "\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace> = \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>" by simp
also have "... = (\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> inv \<lbrace>Cod t\<^bold>\<down>\<rbrace>) \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>"
using comp_assoc by simp
also have "... = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>"
proof -
have "\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> inv \<lbrace>Cod t\<^bold>\<down>\<rbrace> = cod \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
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 \<longleftrightarrow> coherent (Inv t)"
proof
have 1: "\<And>t. Can t \<Longrightarrow> coherent t \<Longrightarrow> coherent (Inv t)"
proof -
fix t
assume "Can t"
hence t: "Can t \<and> Arr t \<and> Ide (Dom t) \<and> Ide (Cod t) \<and>
arr \<lbrace>t\<rbrace> \<and> iso \<lbrace>t\<rbrace> \<and> inverse_arrows \<lbrace>t\<rbrace> (inv \<lbrace>t\<rbrace>) \<and>
Can \<^bold>\<lfloor>t\<^bold>\<rfloor> \<and> Arr \<^bold>\<lfloor>t\<^bold>\<rfloor> \<and> arr \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<and> iso \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<and> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<in> VHom \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<and>
inverse_arrows \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> (inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>) \<and> Inv t \<in> 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 "\<lbrace>Cod (Inv t)\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>Inv t\<rbrace> = (inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>) \<cdot> \<lbrace>Cod (Inv t)\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>Inv t\<rbrace>"
using t comp_inv_arr red_in_Hom
comp_cod_arr [of "\<lbrace>Cod (Inv t)\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>Inv t\<rbrace>" "inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"]
by (auto simp add: eval_simps')
also have "... = inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace> \<cdot> inv \<lbrace>t\<rbrace>"
using t eval_Inv_Can comp_assoc by auto
also have "... = inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom t\<^bold>\<down>\<rbrace>) \<cdot> inv \<lbrace>t\<rbrace>"
using comp_assoc by simp
also have "... = inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace>) \<cdot> inv \<lbrace>t\<rbrace>"
using t coh by simp
also have "... = inv \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace> \<cdot> inv \<lbrace>t\<rbrace>"
using comp_assoc by simp
also have "... = \<lbrace>\<^bold>\<lfloor>Inv t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom (Inv t)\<^bold>\<down>\<rbrace>"
proof -
have "\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace> \<cdot> inv \<lbrace>t\<rbrace> = \<lbrace>Dom (Inv t)\<^bold>\<down>\<rbrace>"
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 \<Longrightarrow> coherent (Inv t)" using assms 1 by simp
show "coherent (Inv t) \<Longrightarrow> 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 \<open>
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.
\<close>
lemma coherent_Prim\<^sub>0:
assumes "C.obj a"
shows "coherent \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0"
by simp
lemma coherent_Prim:
assumes "Arr \<^bold>\<langle>f\<^bold>\<rangle>"
shows "coherent \<^bold>\<langle>f\<^bold>\<rangle>"
using assms by simp
lemma coherent_Lunit_Ide:
assumes "Ide t"
shows "coherent \<^bold>\<l>\<^bold>[t\<^bold>]"
proof -
have t: "Ide t \<and> Arr t \<and> Dom t = t \<and> Cod t = t \<and>
ide \<lbrace>t\<rbrace> \<and> ide \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<and> \<lbrace>t\<^bold>\<down>\<rbrace> \<in> hom \<lbrace>t\<rbrace> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
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 "\<lbrace>Cod \<^bold>\<l>\<^bold>[t\<^bold>]\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace> = \<l>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>] \<cdot> (\<lbrace>Trg t\<rbrace> \<star> \<lbrace>t\<^bold>\<down>\<rbrace>)"
using t \<ll>_ide_simp lunit_naturality [of "\<lbrace>t\<^bold>\<down>\<rbrace>"] red_in_Hom
by (simp add: eval_simps')
also have "... = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<l>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>] \<cdot> (\<lbrace>Trg t\<rbrace> \<star> \<lbrace>t\<^bold>\<down>\<rbrace>)"
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 "... = \<lbrace>\<^bold>\<lfloor>\<^bold>\<l>\<^bold>[t\<^bold>]\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom \<^bold>\<l>\<^bold>[t\<^bold>]\<^bold>\<down>\<rbrace>"
using 1 t Nml_Trg \<ll>_ide_simp by (cases "Trg t"; simp)
finally show ?thesis by simp
qed
text \<open>
Unlike many of the other results, the next one was not quite so straightforward to adapt
from @{session \<open>MonoidalCategory\<close>}.
\<close>
lemma coherent_Runit_Ide:
assumes "Ide t"
shows "coherent \<^bold>\<r>\<^bold>[t\<^bold>]"
proof -
have t: "Ide t \<and> Arr t \<and> Dom t = t \<and> Cod t = t \<and>
ide \<lbrace>t\<rbrace> \<and> ide \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<and> \<lbrace>t\<^bold>\<down>\<rbrace> \<in> hom \<lbrace>t\<rbrace> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
using assms Ide_in_Hom Ide_Nmlize_Ide
red_in_Hom eval_in_hom ide_eval_Ide
by force
have "\<lbrace>Cod \<^bold>\<r>\<^bold>[t\<^bold>]\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace> = \<r>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>] \<cdot> (\<lbrace>t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Src t\<rbrace>)"
using t \<rr>_ide_simp red_in_Hom runit_naturality [of "\<lbrace>t\<^bold>\<down>\<rbrace>"]
by (simp add: eval_simps')
also have "... = \<lbrace>\<^bold>\<lfloor>\<^bold>\<r>\<^bold>[t\<^bold>]\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom \<^bold>\<r>\<^bold>[t\<^bold>]\<^bold>\<down>\<rbrace>"
proof -
have "\<lbrace>\<^bold>\<lfloor>\<^bold>\<r>\<^bold>[t\<^bold>]\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom \<^bold>\<r>\<^bold>[t\<^bold>]\<^bold>\<down>\<rbrace> = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Src t\<^bold>\<down>\<rbrace>)"
using t by (cases t; simp; cases "Src t"; simp)
also have "... = (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<rbrace>) \<cdot> (\<lbrace>t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Src t\<^bold>\<down>\<rbrace>)"
proof -
have "\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<in> hom \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
using t Nmlize_in_Hom by auto
moreover have "\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<rbrace> \<in> hom (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>Src t\<rbrace>) \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
proof -
have "\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<rbrace> \<in> hom \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<rbrace> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
proof -
have "Src \<^bold>\<lfloor>t\<^bold>\<rfloor> = Trg \<^bold>\<lfloor>Src t\<^bold>\<rfloor> \<and> \<^bold>\<lfloor>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<star> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using t Nmlize_Src Nml_Nmlize HcompNml_Nml_Src [of "\<^bold>\<lfloor>t\<^bold>\<rfloor>"]
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 "\<lbrace>t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Src t\<^bold>\<down>\<rbrace> \<in> hom (\<lbrace>t\<rbrace> \<star> \<lbrace>Src t\<rbrace>) (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>Src t\<rbrace>)"
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 "... = \<r>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>] \<cdot> (\<lbrace>t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Src t\<^bold>\<down>\<rbrace>)"
proof -
have "\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Src t\<^bold>\<rfloor>\<rbrace> = \<r>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>]"
proof -
have "Nml \<^bold>\<lfloor>t\<^bold>\<rfloor>" using t Nml_Nmlize by blast
moreover have "is_Prim\<^sub>0 \<^bold>\<lfloor>Src t\<^bold>\<rfloor>"
using t is_Prim0_Src Nmlize_Src by presburger
ultimately show ?thesis
apply (cases "\<^bold>\<lfloor>t\<^bold>\<rfloor>"; simp; cases "\<^bold>\<lfloor>Src t\<^bold>\<rfloor>"; simp)
using t unitor_coincidence \<ll>_ide_simp \<rr>_ide_simp Nmlize_in_Hom
apply simp_all
using t is_Prim0_Src
apply (cases "\<^bold>\<lfloor>t\<^bold>\<rfloor>"; simp)
using t Nmlize_Src unitor_coincidence preserves_obj by simp
qed
moreover have "\<r>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>] \<in> hom (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>Src t\<rbrace>) \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
using t Nmlize_in_Hom by (auto simp add: eval_simps'(2))
ultimately show ?thesis
using comp_cod_arr [of "\<r>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>]"] by fastforce
qed
also have "... = \<r>[\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>] \<cdot> (\<lbrace>t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Src t\<rbrace>)"
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>\<l>\<^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>\<r>\<^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 "\<lbrace>t \<^bold>\<Down> Src t\<rbrace> = \<r>[\<lbrace>t\<rbrace>]"
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 "\<lbrace>Trg t \<^bold>\<Down> t\<rbrace> = \<l>[\<lbrace>t\<rbrace>]"
using assms is_Prim0_Trg [of t] \<ll>_ide_simp ide_eval_Ide
by (cases "Trg t"; simp)
lemma coherence_key_fact:
assumes "Ide a \<and> Nml a" and "Ide b \<and> Nml b" and "Ide c \<and> Nml c"
and "Src a = Trg b" and "Src b = Trg c"
shows "\<lbrace>(a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>a \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>) =
(\<lbrace>a \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "is_Prim\<^sub>0 b \<Longrightarrow> ?thesis"
proof -
assume b: "is_Prim\<^sub>0 b"
have "\<lbrace>a \<^bold>\<Down> c\<rbrace> \<cdot> (\<r>[\<lbrace>a\<rbrace>] \<star> \<lbrace>c\<rbrace>) = (\<lbrace>a \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<l>[\<lbrace>c\<rbrace>])) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>Trg c\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "Src b = Trg b"
using b by (cases b; simp)
thus ?thesis
using assms triangle [of "\<lbrace>c\<rbrace>" "\<lbrace>a\<rbrace>"] 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 "\<lbrakk> \<not> is_Prim\<^sub>0 b; is_Prim\<^sub>0 c \<rbrakk> \<Longrightarrow> ?thesis"
proof -
assume b: "\<not> is_Prim\<^sub>0 b"
assume c: "is_Prim\<^sub>0 c"
have "\<lbrace>(a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>a \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>) = \<lbrace>(a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> Src b\<rbrace> \<cdot> (\<lbrace>a \<^bold>\<Down> b\<rbrace> \<star> src \<lbrace>b\<rbrace>)"
using assms b c by (cases c, simp_all add: eval_simps')
also have "... = \<r>[\<lbrace>a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>] \<cdot> (\<lbrace>a \<^bold>\<Down> b\<rbrace> \<star> src \<lbrace>b\<rbrace>)"
using assms red2_Nml_Src [of "a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b"] Nml_HcompNml(1) Src_HcompNml Ide_HcompNml
by simp
also have "... = \<lbrace>a \<^bold>\<Down> b\<rbrace> \<cdot> \<r>[\<lbrace>a\<rbrace> \<star> \<lbrace>b\<rbrace>]"
proof -
have "\<guillemotleft>\<lbrace>a \<^bold>\<Down> b\<rbrace> : \<lbrace>a\<rbrace> \<star> \<lbrace>b\<rbrace> \<Rightarrow> \<lbrace>a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>\<guillemotright>"
using assms red2_in_Hom eval_in_hom [of "a \<^bold>\<Down> b"] by simp
thus ?thesis
using assms runit_naturality [of "\<lbrace>a \<^bold>\<Down> b\<rbrace>"] arr_dom in_homE src_dom src_hcomp
hcomp_simps(1)
by (elim in_homE, metis)
qed
also have "... = (\<lbrace>a \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "(\<lbrace>a \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>] =
(\<lbrace>a \<^bold>\<Down> b\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<r>[\<lbrace>b\<rbrace>])) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, src \<lbrace>b\<rbrace>]"
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 "\<lbrakk> \<not> is_Prim\<^sub>0 b; \<not> is_Prim\<^sub>0 c \<rbrakk> \<Longrightarrow> ?thesis"
proof -
assume b': "\<not> is_Prim\<^sub>0 b"
hence b: "Ide b \<and> Nml b \<and> Arr b \<and> \<not> is_Prim\<^sub>0 b \<and>
ide \<lbrace>b\<rbrace> \<and> arr \<lbrace>b\<rbrace> \<and> \<^bold>\<lfloor>b\<^bold>\<rfloor> = b \<and> b\<^bold>\<down> = b \<and> Dom b = b \<and> Cod b = b"
using assms Ide_Nmlize_Ide Ide_in_Hom ide_eval_Ide by simp
assume c': "\<not> is_Prim\<^sub>0 c"
hence c: "Ide c \<and> Nml c \<and> Arr c \<and> \<not> is_Prim\<^sub>0 c \<and>
ide \<lbrace>c\<rbrace> \<and> arr \<lbrace>c\<rbrace> \<and> \<^bold>\<lfloor>c\<^bold>\<rfloor> = c \<and> c\<^bold>\<down> = c \<and> Dom c = c \<and> Cod c = c"
using assms Ide_Nmlize_Ide Ide_in_Hom ide_eval_Ide by simp
have "\<And>a. Ide a \<and> Nml a \<and> Src a = Trg b \<Longrightarrow>
\<lbrace>(a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>a \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)
= (\<lbrace>a \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
fix a :: "'c term"
show "Ide a \<and> Nml a \<and> Src a = Trg b \<Longrightarrow>
\<lbrace>(a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>a \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)
= (\<lbrace>a \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> (\<lbrace>a\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
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 \<and> C.arr f \<and> \<^bold>\<langle>src\<^sub>C f\<^bold>\<rangle>\<^sub>0 = Trg b"
show "\<lbrace>(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>) =
(\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "\<lbrace>(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>) =
((E f \<star> \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
((E f \<star> \<lbrace>b\<rbrace>) \<star> \<lbrace>c\<rbrace>)"
proof -
have "((E f \<star> \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
((E f \<star> \<lbrace>b\<rbrace>) \<star> \<lbrace>c\<rbrace>) =
((E f \<star> \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
(\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)"
using f b red2_Nml by simp
also have "... = (\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
(\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)"
proof -
have "\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> = E f \<star> \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>"
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 "... = \<lbrace>(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)"
using b c \<alpha>_def by (cases c, simp_all)
finally show ?thesis by argo
qed
also have "... = ((E f \<star> \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "src (E f) = trg \<lbrace>b\<rbrace>"
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 "... = (\<lbrace>\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> (E f \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[E f, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using assms f b c Ide_HcompNml Nml_HcompNml is_Hcomp_HcompNml [of b c] \<alpha>_def
by (cases "b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c") simp_all
finally show ?thesis by blast
qed
next
fix x
assume x: "C.obj x \<and> \<^bold>\<langle>x\<^bold>\<rangle>\<^sub>0 = Trg b"
show "\<lbrace>b \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>Trg b \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>) =
(\<lbrace>Trg b \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> \<cdot> (\<lbrace>Trg b\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>Trg b\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have 1: "Trg (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c) = Trg b"
using assms b c Trg_HcompNml by blast
have 2: "\<lbrace>Trg b \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> = \<l>[\<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>]"
using assms b c 1 Nml_HcompNml red2_Trg_Nml [of "b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c"] Ide_HcompNml
by simp
have "\<lbrace>b \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>Trg b \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>) = \<lbrace>b \<^bold>\<Down> c\<rbrace> \<cdot> (\<l>[\<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
using b c 1 2 HcompNml_Trg_Nml red2_Trg_Nml Trg_HcompNml by simp
also have "... = \<lbrace>b \<^bold>\<Down> c\<rbrace> \<cdot> \<l>[\<lbrace>b\<rbrace> \<star> \<lbrace>c\<rbrace>] \<cdot> \<a>[\<lbrace>Trg b\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using assms b c lunit_hcomp [of "\<lbrace>b\<rbrace>" "\<lbrace>c\<rbrace>"]
by (metis (no_types, lifting) eval_simps'(3) eval_simps(2))
also have "... = (\<lbrace>b \<^bold>\<Down> c\<rbrace> \<cdot> \<l>[\<lbrace>b\<rbrace> \<star> \<lbrace>c\<rbrace>]) \<cdot> \<a>[\<lbrace>Trg b\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using comp_assoc by simp
also have "... = (\<l>[\<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>] \<cdot> (\<lbrace>Trg b\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>Trg b\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using assms b c lunit_naturality [of "\<lbrace>b \<^bold>\<Down> c\<rbrace>"] red2_in_Hom
by (simp add: eval_simps')
also have "... = (\<lbrace>Trg b \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> \<cdot> (\<lbrace>Trg b\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>Trg b\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
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 \<Longrightarrow> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>e \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)
= (\<lbrace>e \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace> \<cdot> (\<lbrace>e\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
assume de: "Ide d \<and> Ide e \<and> Src d = Trg e \<and> Nml (d \<^bold>\<star> e) \<and> Src e = Trg b"
show "\<lbrace>((d \<^bold>\<star> e) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)
= (\<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> ((\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>) \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
let ?f = "un_Prim d"
have "is_Prim d"
using de Nml_HcompD
by (metis term.disc(12))
hence "d = \<^bold>\<langle>?f\<^bold>\<rangle> \<and> C.ide ?f"
using de by (cases d; simp)
hence d: "Ide d \<and> Arr d \<and> Dom d = d \<and> Cod d = d \<and> Nml d \<and>
d = \<^bold>\<langle>?f\<^bold>\<rangle> \<and> C.ide ?f \<and> ide \<lbrace>d\<rbrace> \<and> arr \<lbrace>d\<rbrace>"
using de ide_eval_Ide Nml_Nmlize(1) Ide_in_Hom Nml_HcompD [of d e]
by simp
have "Nml e \<and> \<not> is_Prim\<^sub>0 e"
using de Nml_HcompD by metis
hence e: "Ide e \<and> Arr e \<and> Dom e = e \<and> Cod e = e \<and> Nml e \<and>
\<not> is_Prim\<^sub>0 e \<and> ide \<lbrace>e\<rbrace> \<and> arr \<lbrace>e\<rbrace>"
using de Ide_in_Hom ide_eval_Ide by simp
have 1: "is_Hcomp (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<and> is_Hcomp (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c) \<and> is_Hcomp (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)"
using assms b c e eb de 1 is_Hcomp_HcompNml [of e b]
is_Hcomp_HcompNml [of "e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b" c] is_Hcomp_HcompNml [of e b]
Nml_HcompNml(1) [of e b] Src_HcompNml
by auto
have "\<lbrace>((d \<^bold>\<star> e) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace> \<cdot> (\<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)
= ((\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>) \<cdot> \<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
proof -
have "\<lbrace>((d \<^bold>\<star> e) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>
= (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "((d \<^bold>\<star> e) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c = (d \<^bold>\<star> (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b)) \<^bold>\<Down> 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>\<Down> ((e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)) \<^bold>\<cdot> (d \<^bold>\<star> ((e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c)) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[d, e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b, c\<^bold>]"
using b c d e de 1 Nml_HcompNml Nmlize_Nml
by (cases c) simp_all
also have "... = (d \<^bold>\<star> ((e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)) \<^bold>\<cdot> (d \<^bold>\<star> ((e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c)) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[d, e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b, c\<^bold>]"
using d 4
apply (cases d, simp_all)
by (cases "(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c") simp_all
finally show ?thesis
using b c d e HcompNml_in_Hom red2_in_Hom
Nml_HcompNml Ide_HcompNml \<alpha>_def
by simp
qed
moreover have "\<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> b\<rbrace>
= (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>) \<cdot> \<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>]"
proof -
have "(d \<^bold>\<star> e) \<^bold>\<Down> b = (d \<^bold>\<Down> (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b)) \<^bold>\<cdot> (d \<^bold>\<star> (e \<^bold>\<Down> b)) \<^bold>\<cdot> \<^bold>\<a>\<^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>\<star> (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b)) \<^bold>\<cdot> (d \<^bold>\<star> (e \<^bold>\<Down> b)) \<^bold>\<cdot> \<^bold>\<a>\<^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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b", simp_all)
finally show ?thesis
using b d e HcompNml_in_Hom red2_in_Hom \<alpha>_def by simp
qed
ultimately show ?thesis by argo
qed
also have "... = ((\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>) \<star> \<lbrace>c\<rbrace>) \<cdot> (\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
proof -
have "(\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>) \<cdot> \<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>
= ((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>) \<star> \<lbrace>c\<rbrace>) \<cdot> (\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
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 "\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>" "\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>]" "\<lbrace>c\<rbrace>" "\<lbrace>c\<rbrace>"]
by (auto simp add: eval_simps')
moreover have "(\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>] =
(\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "(\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>] =
((\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>)) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>]"
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 "... = (\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> (\<lbrace>e \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace> \<star> \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>] \<cdot> (\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
using assms b c d e de HcompNml_in_Hom red2_in_Hom
Ide_HcompNml Nml_HcompNml ide_eval_Ide
assoc_naturality [of "\<lbrace>d\<rbrace>" "\<lbrace>e \<^bold>\<Down> b\<rbrace>" "\<lbrace>c\<rbrace>"]
comp_permute [of "\<a>[\<lbrace>d\<rbrace>, \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b\<rbrace>, \<lbrace>c\<rbrace>]" "(\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b\<rbrace>) \<star> \<lbrace>c\<rbrace>"
"\<lbrace>d\<rbrace> \<star> (\<lbrace>e \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>)" "\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace> \<star> \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"]
comp_assoc
by (simp add: eval_simps')
also have "... = ((\<lbrace>d\<rbrace> \<star> \<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> (\<lbrace>e \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>))) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace> \<star> \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>] \<cdot> (\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
using comp_assoc by simp
also have "... = (((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot>
(\<lbrace>d\<rbrace> \<star> \<a>[\<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>])) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace> \<star> \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>] \<cdot> (\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>)"
using assms b c d e de eb I HcompNml_in_Hom red2_in_Hom
Ide_HcompNml Nml_HcompNml whisker_left [of "\<lbrace>d\<rbrace>"]
interchange [of "\<lbrace>d\<rbrace>" "\<lbrace>d\<rbrace>" "\<lbrace>(e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b) \<^bold>\<Down> c\<rbrace>" "\<lbrace>e \<^bold>\<Down> b\<rbrace> \<star> \<lbrace>c\<rbrace>"]
by (auto simp add: eval_simps')
also have "... = ((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot>
((\<lbrace>d\<rbrace> \<star> \<a>[\<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace> \<star> \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>] \<cdot> (\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>] \<star> \<lbrace>c\<rbrace>))"
using comp_assoc by simp
also have "... = ((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> (\<lbrace>e\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>))) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace> \<star> \<lbrace>c\<rbrace>] \<cdot> \<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using assms b c d e de pentagon
by (simp add: eval_simps')
also have "... = (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>) \<cdot>
((\<lbrace>d\<rbrace> \<star> (\<lbrace>e\<rbrace> \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot> \<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace> \<star> \<lbrace>c\<rbrace>]) \<cdot>
\<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using comp_assoc by simp
also have "... = ((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>)) \<cdot>
(\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>] \<cdot> ((\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>) \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot>
\<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using assms d e de HcompNml_in_Hom red2_in_Hom Ide_HcompNml Nml_HcompNml
assoc_naturality [of "\<lbrace>d\<rbrace>" "\<lbrace>e\<rbrace>" "\<lbrace>b \<^bold>\<Down> c\<rbrace>"] comp_cod_arr [of "\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>"]
by (simp add: eval_simps')
also have "... = ((\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>]) \<cdot>
((\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>) \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>) \<cdot> \<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using comp_assoc by simp
also have "... = \<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> ((\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>) \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
proof -
have "\<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>
= (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>) \<cdot> (\<lbrace>d\<rbrace> \<star> \<lbrace>e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace>) \<cdot>
\<a>[\<lbrace>d\<rbrace>, \<lbrace>e\<rbrace>, \<lbrace>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<rbrace>]"
proof -
have "(d \<^bold>\<star> e) \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)
= (d \<^bold>\<Down> (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<^bold>\<rfloor>)) \<^bold>\<cdot> (d \<^bold>\<star> (e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c))) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[d, e, b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<^bold>]"
using e 1 by (cases "b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c") auto
also have "... = (d \<^bold>\<Down> (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c))) \<^bold>\<cdot> (d \<^bold>\<star> (e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c))) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[d, e, b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<^bold>]"
using assms Nml_HcompNml Nmlize_Nml by simp
also have "... = (d \<^bold>\<star> (e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c))) \<^bold>\<cdot> (d \<^bold>\<star> (e \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c))) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[d, e, b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c\<^bold>]"
using d 1
apply (cases "e \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c", simp_all)
by (cases d, simp_all)
finally show ?thesis
using \<alpha>_def by simp
qed
thus ?thesis by simp
qed
also have "... = (\<lbrace>(d \<^bold>\<star> e) \<^bold>\<Down> (b \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> c)\<rbrace> \<cdot> ((\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>) \<star> \<lbrace>b \<^bold>\<Down> c\<rbrace>)) \<cdot>
\<a>[\<lbrace>d\<rbrace> \<star> \<lbrace>e\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
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>\<a>\<^bold>[a, b, c\<^bold>]"
proof -
have a: "Ide a \<and> Arr a \<and> Dom a = a \<and> Cod a = a \<and>
ide \<lbrace>a\<rbrace> \<and> ide \<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor>\<rbrace> \<and> \<lbrace>a\<^bold>\<down>\<rbrace> \<in> hom \<lbrace>a\<rbrace> \<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor>\<rbrace>"
using assms Ide_in_Hom Ide_Nmlize_Ide ide_eval_Ide red_in_Hom eval_in_hom(2)
by force
have b: "Ide b \<and> Arr b \<and> Dom b = b \<and> Cod b = b \<and>
ide \<lbrace>b\<rbrace> \<and> ide \<lbrace>\<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<and> \<lbrace>b\<^bold>\<down>\<rbrace> \<in> hom \<lbrace>b\<rbrace> \<lbrace>\<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace>"
using assms Ide_in_Hom Ide_Nmlize_Ide ide_eval_Ide red_in_Hom(2)
eval_in_hom(2) [of "b\<^bold>\<down>"]
by auto
have c: "Ide c \<and> Arr c \<and> Dom c = c \<and> Cod c = c \<and>
ide \<lbrace>c\<rbrace> \<and> ide \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> \<and> \<lbrace>c\<^bold>\<down>\<rbrace> \<in> hom \<lbrace>c\<rbrace> \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>"
using assms Ide_in_Hom Ide_Nmlize_Ide red_in_Hom eval_in_hom(2) [of "c\<^bold>\<down>"]
ide_eval_Ide
by auto
have "\<lbrace>Cod \<^bold>\<a>\<^bold>[a, b, c\<^bold>]\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>\<^bold>\<a>\<^bold>[a, b, c\<^bold>]\<rbrace>
= (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> (\<^bold>\<lfloor>b\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>c\<^bold>\<rfloor>)\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>b\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>) \<cdot> (\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace> \<star> \<lbrace>c\<^bold>\<down>\<rbrace>)) \<cdot>
\<a>[\<lbrace>a\<rbrace>, \<lbrace>b\<rbrace>, \<lbrace>c\<rbrace>]"
using assms a b c red_in_Hom red2_in_Hom Nml_Nmlize Ide_Nmlize_Ide
\<alpha>_def eval_red_Hcomp interchange [of "\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor>\<rbrace>" "\<lbrace>a\<^bold>\<down>\<rbrace>"] comp_cod_arr [of "\<lbrace>a\<^bold>\<down>\<rbrace>"]
by (simp add: eval_simps')
also have "... = ((\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> (\<^bold>\<lfloor>b\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>c\<^bold>\<rfloor>)\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>b\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>)) \<cdot> \<a>[\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor>\<rbrace>, \<lbrace>\<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace>, \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>]) \<cdot>
((\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>) \<star> \<lbrace>c\<^bold>\<down>\<rbrace>)"
using assms red_in_Hom Ide_HcompNml assoc_naturality [of "\<lbrace>a\<^bold>\<down>\<rbrace>" "\<lbrace>b\<^bold>\<down>\<rbrace>" "\<lbrace>c\<^bold>\<down>\<rbrace>"] comp_assoc
by (simp add: eval_simps')
also have "... = (\<lbrace>(\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor>) \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>)) \<cdot> ((\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>) \<star> \<lbrace>c\<^bold>\<down>\<rbrace>)"
using assms Nml_Nmlize Ide_Nmlize_Ide coherence_key_fact by simp
also have "... = \<lbrace>\<^bold>\<lfloor>\<^bold>\<a>\<^bold>[a, b, c\<^bold>]\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom \<^bold>\<a>\<^bold>[a, b, c\<^bold>]\<^bold>\<down>\<rbrace>"
using assms a b c red_in_Hom red2_in_Hom Ide_Nmlize_Ide
Nml_Nmlize eval_red_Hcomp HcompNml_assoc
interchange [of "\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace>" "\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>" "\<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>" "\<lbrace>c\<^bold>\<down>\<rbrace>"]
comp_cod_arr [of "\<lbrace>c\<^bold>\<down>\<rbrace>" "\<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>"]
apply (simp add: eval_simps')
proof -
have "seq \<lbrace>(\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor>) \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> ((\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>) \<cdot> ((\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>) \<star> \<lbrace>c\<^bold>\<down>\<rbrace>))"
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 (\<lbrace>(\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor>) \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>) \<cdot> ((\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>) \<star> \<lbrace>c\<^bold>\<down>\<rbrace>)) =
\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>"
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 "(\<lbrace>(\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor>) \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>)) \<cdot> ((\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>) \<star> \<lbrace>c\<^bold>\<down>\<rbrace>) =
\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> \<cdot>
\<lbrace>(\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>b\<^bold>\<rfloor>) \<^bold>\<Down> \<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>a\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>b\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>c\<^bold>\<rfloor>\<rbrace>) \<cdot> ((\<lbrace>a\<^bold>\<down>\<rbrace> \<star> \<lbrace>b\<^bold>\<down>\<rbrace>) \<star> \<lbrace>c\<^bold>\<down>\<rbrace>)"
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>\<a>\<^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 "\<lbrace>Cod t \<^bold>\<Down> Cod u\<rbrace> \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>) = \<lbrace>t \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>Dom t \<^bold>\<Down> Dom u\<rbrace>"
proof -
have *: "\<And>t u. Nml (t \<^bold>\<star> u) \<Longrightarrow> arr \<lbrace>t\<rbrace> \<and> arr \<lbrace>u\<rbrace>"
using Nml_implies_Arr Nml_HcompD
by (metis eval_simps'(1))
have "is_Prim\<^sub>0 t \<Longrightarrow> ?thesis"
using assms Nml_implies_Arr is_Prim0_Trg \<ll>.naturality [of "\<lbrace>u\<rbrace>"]
by (cases t) (simp_all add: eval_simps', cases "Trg t", simp_all)
moreover have "\<not> is_Prim\<^sub>0 t \<and> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
using assms Nml_implies_Arr eval_red2_Nml_Prim\<^sub>0 runit_naturality [of "\<lbrace>t\<rbrace>"]
by (cases u) (simp_all add: eval_simps')
moreover have "\<not> is_Prim\<^sub>0 t \<and> \<not> is_Prim\<^sub>0 u \<Longrightarrow> ?thesis"
using assms * Nml_implies_Arr
apply (induct t, simp_all)
proof -
fix f
assume f: "C.arr f"
assume "\<not> is_Prim\<^sub>0 u"
hence u: "\<not> is_Prim\<^sub>0 u \<and>
Nml u \<and> Nml (Dom u) \<and> Nml (Cod u) \<and> Ide (Dom u) \<and> Ide (Cod u) \<and>
arr \<lbrace>u\<rbrace> \<and> arr \<lbrace>Dom u\<rbrace> \<and> arr \<lbrace>Cod u\<rbrace> \<and> ide \<lbrace>Dom u\<rbrace> \<and> ide \<lbrace>Cod u\<rbrace>"
using assms(2) Nml_implies_Arr ide_eval_Ide by simp
hence 1: "\<not> is_Prim\<^sub>0 (Dom u) \<and> \<not> is_Prim\<^sub>0 (Cod u)"
using u by (cases u) simp_all
assume "\<^bold>\<langle>src\<^sub>C f\<^bold>\<rangle>\<^sub>0 = Trg u"
hence "\<lbrace>\<^bold>\<langle>src\<^sub>C f\<^bold>\<rangle>\<^sub>0\<rbrace> = \<lbrace>Trg u\<rbrace>" by simp
hence fu: "src (E f) = trg \<lbrace>u\<rbrace>"
using f u preserves_src Nml_implies_Arr
by (simp add: eval_simps')
show "\<lbrace>\<^bold>\<langle>C.cod f\<^bold>\<rangle> \<^bold>\<Down> Cod u\<rbrace> \<cdot> (E f \<star> \<lbrace>u\<rbrace>) = (E f \<star> \<lbrace>u\<rbrace>) \<cdot> \<lbrace>\<^bold>\<langle>C.dom f\<^bold>\<rangle> \<^bold>\<Down> Dom u\<rbrace>"
proof -
have "\<lbrace>\<^bold>\<langle>C.cod f\<^bold>\<rangle> \<^bold>\<Down> Cod u\<rbrace> = E (C.cod f) \<star> cod \<lbrace>u\<rbrace>"
using f u 1 Nml_implies_Arr
by (cases "Cod u", simp_all add: eval_simps')
moreover have "\<lbrace>\<^bold>\<langle>C.dom f\<^bold>\<rangle> \<^bold>\<Down> Dom u\<rbrace> = E (C.dom f) \<star> dom \<lbrace>u\<rbrace>"
using f u 1 Nml_implies_Arr
by (cases "Dom u", simp_all add: eval_simps')
moreover have "\<guillemotleft>E f \<star> \<lbrace>u\<rbrace> : E (C.dom f) \<star> \<lbrace>Dom u\<rbrace> \<Rightarrow> E (C.cod f) \<star> \<lbrace>Cod u\<rbrace>\<guillemotright>"
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: "\<lbrakk> \<not> is_Prim\<^sub>0 w; Nml w \<rbrakk> \<Longrightarrow>
\<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace> \<cdot> (\<lbrace>w\<rbrace> \<star> \<lbrace>u\<rbrace>) = \<lbrace>w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>"
assume "\<not> is_Prim\<^sub>0 u"
hence u: "\<not> is_Prim\<^sub>0 u \<and> Arr u \<and> Arr (Dom u) \<and> Arr (Cod u) \<and>
Nml u \<and> Nml (Dom u) \<and> Nml (Cod u) \<and> Ide (Dom u) \<and> Ide (Cod u) \<and>
arr \<lbrace>u\<rbrace> \<and> arr \<lbrace>Dom u\<rbrace> \<and> arr \<lbrace>Cod u\<rbrace> \<and> ide \<lbrace>Dom u\<rbrace> \<and> ide \<lbrace>Cod u\<rbrace>"
using assms(2) Nml_implies_Arr ide_eval_Ide by simp
assume vw: "Nml (v \<^bold>\<star> w)"
assume wu: "Src w = Trg u"
let ?f = "un_Prim v"
have "v = \<^bold>\<langle>?f\<^bold>\<rangle> \<and> C.arr ?f"
using vw by (metis Nml_HcompD(1) Nml_HcompD(2))
hence "Arr v \<and> v = \<^bold>\<langle>un_Prim v\<^bold>\<rangle> \<and> C.arr ?f \<and> Nml v" by (cases v; simp)
hence v: "v = \<^bold>\<langle>?f\<^bold>\<rangle> \<and> C.arr ?f \<and> Arr v \<and> Arr (Dom v) \<and> Arr (Cod v) \<and> Nml v \<and>
Nml (Dom v) \<and> Nml (Cod v) \<and>
arr \<lbrace>v\<rbrace> \<and> arr \<lbrace>Dom v\<rbrace> \<and> arr \<lbrace>Cod v\<rbrace> \<and> ide \<lbrace>Dom v\<rbrace> \<and> ide \<lbrace>Cod v\<rbrace>"
using vw * by (cases v, simp_all)
have "Nml w \<and> \<not> is_Prim\<^sub>0 w"
using vw v by (metis Nml.simps(3))
hence w: "\<not> is_Prim\<^sub>0 w \<and> Arr w \<and> Arr (Dom w) \<and> Arr (Cod w) \<and>
Nml w \<and> Nml (Dom w) \<and> Nml (Cod w) \<and>
Ide (Dom w) \<and> Ide (Cod w) \<and>
arr \<lbrace>w\<rbrace> \<and> arr \<lbrace>Dom w\<rbrace> \<and> arr \<lbrace>Cod w\<rbrace> \<and> ide \<lbrace>Dom w\<rbrace> \<and> ide \<lbrace>Cod w\<rbrace>"
using vw * Nml_implies_Arr ide_eval_Ide by simp
have u': "\<not> is_Prim\<^sub>0 (Dom u) \<and> \<not> is_Prim\<^sub>0 (Cod u)"
using u by (cases u, simp_all)
have w': "\<not> is_Prim\<^sub>0 (Dom w) \<and> \<not> 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>\<star> (Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u))"
using u u' v w w' wu vw is_Hcomp_HcompNml Nml_HcompNml Nml_Dom
by (metis Dom.simps(3) Nml.simps(3) term.distinct_disc(3))
have Y: "Nml (Cod v \<^bold>\<star> (Cod w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u))"
using u u' w w' wu vw is_Hcomp_HcompNml Nml_HcompNml Src_Cod Trg_Cod
by (metis Cod.simps(3) Nml.simps(3) Nml_Cod term.distinct_disc(3) v)
show "\<lbrace>(Cod v \<^bold>\<star> Cod w) \<^bold>\<Down> Cod u\<rbrace> \<cdot> ((\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<star> \<lbrace>u\<rbrace>)
= \<lbrace>(v \<^bold>\<star> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>(Dom v \<^bold>\<star> Dom w) \<^bold>\<Down> Dom u\<rbrace>"
proof -
have "\<lbrace>(Cod v \<^bold>\<star> Cod w) \<^bold>\<Down> Cod u\<rbrace> \<cdot> ((\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<star> \<lbrace>u\<rbrace>)
= (\<lbrace>Cod v \<^bold>\<Down> (Cod w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u)\<rbrace> \<cdot> (\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot>
\<a>[\<lbrace>Cod v\<rbrace>, \<lbrace>Cod w\<rbrace>, \<lbrace>Cod u\<rbrace>]) \<cdot> ((\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<star> \<lbrace>u\<rbrace>)"
proof -
have "(Cod v \<^bold>\<star> Cod w) \<^bold>\<Down> Cod u
= (Cod v \<^bold>\<Down> (Cod w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod u\<^bold>\<rfloor>)) \<^bold>\<cdot> (Cod v \<^bold>\<star> Cod w \<^bold>\<Down> Cod u) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[Cod v, Cod w, Cod u\<^bold>]"
using u v w by (cases u) simp_all
hence "\<lbrace>(Cod v \<^bold>\<star> Cod w) \<^bold>\<Down> Cod u\<rbrace>
= \<lbrace>Cod v \<^bold>\<Down> (Cod w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u)\<rbrace> \<cdot> (\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot>
\<a>[\<lbrace>Cod v\<rbrace>, \<lbrace>Cod w\<rbrace>, \<lbrace>Cod u\<rbrace>]"
using u v w \<alpha>_def by simp
thus ?thesis by presburger
qed
also have "... = ((\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod u\<rbrace>) \<cdot> (\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot>
\<a>[\<lbrace>Cod v\<rbrace>, \<lbrace>Cod w\<rbrace>, \<lbrace>Cod u\<rbrace>]) \<cdot> ((\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<star> \<lbrace>u\<rbrace>)"
using u v w Y red2_Nml by simp
also have "... = ((\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot> \<a>[\<lbrace>Cod v\<rbrace>, \<lbrace>Cod w\<rbrace>, \<lbrace>Cod u\<rbrace>]) \<cdot>
((\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<star> \<lbrace>u\<rbrace>)"
using u v w vw' wu comp_cod_arr red2_in_Hom HcompNml_in_Hom comp_reduce
by (simp add: eval_simps')
also have "... = (\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot> \<a>[\<lbrace>Cod v\<rbrace>, \<lbrace>Cod w\<rbrace>, \<lbrace>Cod u\<rbrace>] \<cdot>
((\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace>) \<star> \<lbrace>u\<rbrace>)"
using comp_assoc by simp
also have "... = (\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot> (\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace> \<star> \<lbrace>u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using u v w vw' wu assoc_naturality [of "\<lbrace>v\<rbrace>" "\<lbrace>w\<rbrace>" "\<lbrace>u\<rbrace>"]
by (simp add: eval_simps')
also have "... = ((\<lbrace>Cod v\<rbrace> \<star> \<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>) \<cdot> (\<lbrace>v\<rbrace> \<star> \<lbrace>w\<rbrace> \<star> \<lbrace>u\<rbrace>)) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using comp_assoc by simp
also have
"... = (\<lbrace>v\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot> \<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using v w u vw' wu I2 red2_in_Hom HcompNml_in_Hom comp_cod_arr
interchange [of "\<lbrace>Cod v\<rbrace>" "\<lbrace>v\<rbrace>" "\<lbrace>Cod w \<^bold>\<Down> Cod u\<rbrace>" "\<lbrace>w\<rbrace> \<star> \<lbrace>u\<rbrace>"]
by (simp add: eval_simps')
also have "... = ((\<lbrace>v\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace>) \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>)) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using v w u vw' wu red2_in_Hom HcompNml_in_Hom comp_arr_dom
interchange [of "\<lbrace>v\<rbrace>" "\<lbrace>Dom v\<rbrace>" "\<lbrace>w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace>" "\<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>"]
by (simp add: eval_simps')
also have "... = (\<lbrace>v\<rbrace> \<star> \<lbrace>w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace>) \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using comp_assoc by simp
also have "... = \<lbrace>v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using u u' v w vw' wu is_Hcomp_HcompNml HcompNml_Prim [of "w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u" ?f]
by force
also have "... = \<lbrace>v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u\<rbrace> \<cdot>
(\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot> \<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
proof -
have "\<lbrace>v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>] =
(\<lbrace>v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u\<rbrace>) \<cdot>
(\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot> \<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using u v w vw' wu comp_arr_dom Nml_HcompNml HcompNml_in_Hom
by (simp add: eval_simps')
also have "... = \<lbrace>v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u\<rbrace> \<cdot>
(\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot> \<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using comp_assoc by simp
finally show ?thesis by blast
qed
also have "... = \<lbrace>(v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> w) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> u\<rbrace> \<cdot> \<lbrace>(Dom v \<^bold>\<star> Dom w) \<^bold>\<Down> Dom u\<rbrace>"
proof -
have "(Dom v \<^bold>\<star> Dom w) \<^bold>\<Down> Dom u
= (Dom v \<^bold>\<Down> (Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom u\<^bold>\<rfloor>)) \<^bold>\<cdot> (Dom v \<^bold>\<star> (Dom w \<^bold>\<Down> Dom u)) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[Dom v, Dom w, Dom u\<^bold>]"
using u u' v w vw' wu by (cases u, simp_all)
hence
"\<lbrace>(Dom v \<^bold>\<star> Dom w) \<^bold>\<Down> Dom u\<rbrace>
= \<lbrace>Dom v \<^bold>\<Down> (Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u)\<rbrace> \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using u v w \<alpha>_def by simp
also have
"... = \<lbrace>Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u\<rbrace> \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
using X HcompNml_Nml red2_Nml by presburger
finally have
"\<lbrace>(Dom v \<^bold>\<star> Dom w) \<^bold>\<Down> Dom u\<rbrace>
= \<lbrace>Dom v \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom w \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom u\<rbrace> \<cdot> (\<lbrace>Dom v\<rbrace> \<star> \<lbrace>Dom w \<^bold>\<Down> Dom u\<rbrace>) \<cdot>
\<a>[\<lbrace>Dom v\<rbrace>, \<lbrace>Dom w\<rbrace>, \<lbrace>Dom u\<rbrace>]"
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>\<star> u)"
proof -
have t: "Arr t \<and> Ide (Dom t) \<and> Ide (Cod t) \<and> Ide \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<and> Ide \<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<and>
arr \<lbrace>t\<rbrace> \<and> arr \<lbrace>Dom t\<rbrace> \<and> ide \<lbrace>Dom t\<rbrace> \<and> arr \<lbrace>Cod t\<rbrace> \<and> ide \<lbrace>Cod t\<rbrace>"
using assms Ide_Nmlize_Ide ide_eval_Ide by auto
have u: "Arr u \<and> Ide (Dom u) \<and> Ide (Cod u) \<and> Ide \<^bold>\<lfloor>Dom u\<^bold>\<rfloor> \<and> Ide \<^bold>\<lfloor>Cod u\<^bold>\<rfloor> \<and>
arr \<lbrace>u\<rbrace> \<and> arr \<lbrace>Dom u\<rbrace> \<and> ide \<lbrace>Dom u\<rbrace> \<and> arr \<lbrace>Cod u\<rbrace> \<and> ide \<lbrace>Cod u\<rbrace>"
using assms Ide_Nmlize_Ide ide_eval_Ide by auto
have "\<lbrace>Cod (t \<^bold>\<star> u)\<^bold>\<down>\<rbrace> \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>)
= (\<lbrace>\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Cod u\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Cod u\<^bold>\<down>\<rbrace>)) \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>)"
using t u eval_red_Hcomp by simp
also have "... = \<lbrace>\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Cod u\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Cod u\<^bold>\<down>\<rbrace>) \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>)"
using comp_assoc by simp
also have "... = \<lbrace>\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Cod u\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace>) \<cdot> (\<lbrace>Dom t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Dom u\<^bold>\<down>\<rbrace>)"
using assms t u Nmlize_in_Hom red_in_Hom
interchange [of "\<lbrace>Cod t\<^bold>\<down>\<rbrace>" "\<lbrace>t\<rbrace>" "\<lbrace>Cod u\<^bold>\<down>\<rbrace>" "\<lbrace>u\<rbrace>"]
interchange [of "\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>" "\<lbrace>Dom t\<^bold>\<down>\<rbrace>" "\<lbrace>\<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace>" "\<lbrace>Dom u\<^bold>\<down>\<rbrace>"]
by (simp add: eval_simps')
also have "... = (\<lbrace>\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Cod u\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace> \<star> \<lbrace>\<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace>)) \<cdot> (\<lbrace>Dom t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Dom u\<^bold>\<down>\<rbrace>)"
using comp_assoc by simp
also have "... = (\<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Dom u\<^bold>\<rfloor>\<rbrace>) \<cdot> (\<lbrace>Dom t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Dom u\<^bold>\<down>\<rbrace>)"
using assms t u Nml_Nmlize Nmlize_in_Hom
eval_red2_naturality [of "Nmlize t" "Nmlize u"]
by simp
also have "... = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<Down> \<^bold>\<lfloor>Dom u\<^bold>\<rfloor>\<rbrace> \<cdot> (\<lbrace>Dom t\<^bold>\<down>\<rbrace> \<star> \<lbrace>Dom u\<^bold>\<down>\<rbrace>)"
using comp_assoc by simp
also have "... = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>(Dom t \<^bold>\<star> Dom u)\<^bold>\<down>\<rbrace>"
using t u eval_red_Hcomp by simp
finally have "\<lbrace>Cod (t \<^bold>\<star> u)\<^bold>\<down>\<rbrace> \<cdot> (\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>) = \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>(Dom t \<^bold>\<star> Dom u)\<^bold>\<down>\<rbrace>"
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>\<cdot> u)"
proof -
have t: "Arr t \<and> Ide (Dom t) \<and> Ide (Cod t) \<and> Ide \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<and> Ide \<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<and>
arr \<lbrace>t\<rbrace> \<and> arr \<lbrace>Dom t\<rbrace> \<and> ide \<lbrace>Dom t\<rbrace> \<and> arr \<lbrace>Cod t\<rbrace> \<and> ide \<lbrace>Cod t\<rbrace>"
using assms Ide_Nmlize_Ide ide_eval_Ide by auto
have u: "Arr u \<and> Ide (Dom u) \<and> Ide (Cod u) \<and> Ide \<^bold>\<lfloor>Dom u\<^bold>\<rfloor> \<and> Ide \<^bold>\<lfloor>Cod u\<^bold>\<rfloor> \<and>
arr \<lbrace>u\<rbrace> \<and> arr \<lbrace>Dom u\<rbrace> \<and> ide \<lbrace>Dom u\<rbrace> \<and> arr \<lbrace>Cod u\<rbrace> \<and> ide \<lbrace>Cod u\<rbrace>"
using assms Ide_Nmlize_Ide ide_eval_Ide by auto
have "\<lbrace>Cod (t \<^bold>\<cdot> u)\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t \<^bold>\<cdot> u\<rbrace> = \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace> \<cdot> \<lbrace>u\<rbrace>"
using t u by simp
also have "... = (\<lbrace>Cod t\<^bold>\<down>\<rbrace> \<cdot> \<lbrace>t\<rbrace>) \<cdot> \<lbrace>u\<rbrace>"
proof -
have "seq \<lbrace>Cod t\<^bold>\<down>\<rbrace> \<lbrace>t\<rbrace>"
using assms t red_in_Hom
by (intro seqI, auto simp add: eval_simps')
moreover have "seq \<lbrace>t\<rbrace> \<lbrace>u\<rbrace>"
using assms t u by (auto simp add: eval_simps')
ultimately show ?thesis
using comp_assoc by auto
qed
also have "... = \<lbrace>\<^bold>\<lfloor>t \<^bold>\<cdot> u\<^bold>\<rfloor>\<rbrace> \<cdot> \<lbrace>Dom (t \<^bold>\<cdot> u)\<^bold>\<down>\<rbrace>"
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>\<cdot> u)" by blast
qed
text \<open>
The main result: ``Every formal arrow is coherent.''
\<close>
theorem coherence:
assumes "Arr t"
shows "coherent t"
proof -
have "Arr t \<Longrightarrow> coherent t"
proof (induct t)
show "\<And>a. Arr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<Longrightarrow> coherent \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0" by simp
show "\<And>\<mu>. Arr \<^bold>\<langle>\<mu>\<^bold>\<rangle> \<Longrightarrow> coherent \<^bold>\<langle>\<mu>\<^bold>\<rangle>" by simp
fix u v
show "\<lbrakk> Arr u \<Longrightarrow> coherent u; Arr v \<Longrightarrow> coherent v \<rbrakk> \<Longrightarrow> Arr (u \<^bold>\<star> v)
\<Longrightarrow> coherent (u \<^bold>\<star> v)"
using coherent_Hcomp by simp
show "\<lbrakk> Arr u \<Longrightarrow> coherent u; Arr v \<Longrightarrow> coherent v \<rbrakk> \<Longrightarrow> Arr (u \<^bold>\<cdot> v)
\<Longrightarrow> coherent (u \<^bold>\<cdot> v)"
using coherent_Vcomp by simp
next
fix t
assume I: "Arr t \<Longrightarrow> coherent t"
show Lunit: "Arr \<^bold>\<l>\<^bold>[t\<^bold>] \<Longrightarrow> coherent \<^bold>\<l>\<^bold>[t\<^bold>]"
using I Ide_Dom coherent_Lunit_Ide Ide_in_Hom
coherent_Vcomp [of t "\<^bold>\<l>\<^bold>[Dom t\<^bold>]"] Nmlize_Vcomp_Arr_Dom
eval_in_hom \<ll>.is_natural_1 [of "\<lbrace>t\<rbrace>"]
by force
show Runit: "Arr \<^bold>\<r>\<^bold>[t\<^bold>] \<Longrightarrow> coherent \<^bold>\<r>\<^bold>[t\<^bold>]"
using I Ide_Dom coherent_Runit_Ide Ide_in_Hom ide_eval_Ide
coherent_Vcomp [of t "\<^bold>\<r>\<^bold>[Dom t\<^bold>]"] Nmlize_Vcomp_Arr_Dom \<rr>_ide_simp
eval_in_hom \<rr>.is_natural_1 [of "\<lbrace>t\<rbrace>"]
by force
show "Arr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow> coherent \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]"
proof -
assume "Arr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]"
hence t: "Arr t" by simp
have "coherent (\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\<cdot> t)"
using t I Ide_Cod coherent_Lunit'_Ide Ide_in_Hom coherent_Vcomp [of "\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>]" t]
Arr.simps(6) Dom.simps(6) Dom_Cod Ide_implies_Arr
by presburger
moreover have "\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\<cdot> t\<rbrace>"
using t \<ll>'.is_natural_2 [of "\<lbrace>t\<rbrace>"]
by (simp add: eval_simps(5))
ultimately show ?thesis
using t Nmlize_Vcomp_Cod_Arr by simp
qed
show "Arr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow> coherent \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]"
proof -
assume "Arr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]"
hence t: "Arr t" by simp
have "coherent (\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\<cdot> t)"
using t I Ide_Cod coherent_Runit'_Ide Ide_in_Hom coherent_Vcomp [of "\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>]" t]
Arr.simps(8) Dom.simps(8) Dom_Cod Ide_implies_Arr
by presburger
moreover have "\<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace> = \<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\<cdot> t\<rbrace>"
using t \<rr>'.is_natural_2 [of "\<lbrace>t\<rbrace>"]
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 \<Longrightarrow> coherent t"
assume I2: "Arr u \<Longrightarrow> coherent u"
assume I3: "Arr v \<Longrightarrow> coherent v"
show "Arr \<^bold>\<a>\<^bold>[t, u, v\<^bold>] \<Longrightarrow> coherent \<^bold>\<a>\<^bold>[t, u, v\<^bold>]"
proof -
assume tuv: "Arr \<^bold>\<a>\<^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>\<star> u \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[Dom t, Dom u, Dom v\<^bold>])"
proof -
have "Arr (t \<^bold>\<star> u \<^bold>\<star> v) \<and> coherent (t \<^bold>\<star> u \<^bold>\<star> 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>\<a>\<^bold>[Dom t, Dom u, Dom v\<^bold>]"
using t u v tu uv Ide_Dom by simp
moreover have "coherent \<^bold>\<a>\<^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>\<star> u \<^bold>\<star> v) = Cod \<^bold>\<a>\<^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>\<a>\<^bold>[t, u, v\<^bold>] ((t \<^bold>\<star> u \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[Dom t, Dom u, Dom v\<^bold>])"
using t u v tu uv Ide_Dom by simp
moreover have "\<^bold>\<lfloor>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>(t \<^bold>\<star> u \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<^bold>\<rfloor>"
proof -
have "(\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>
= (\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> ((\<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom u\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom v\<^bold>\<rfloor>)"
proof -
have 1: "Nml \<^bold>\<lfloor>t\<^bold>\<rfloor> \<and> Nml \<^bold>\<lfloor>u\<^bold>\<rfloor> \<and> Nml \<^bold>\<lfloor>v\<^bold>\<rfloor> \<and>
Dom \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> \<and> Dom \<^bold>\<lfloor>u\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom u\<^bold>\<rfloor> \<and> Dom \<^bold>\<lfloor>v\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom v\<^bold>\<rfloor>"
using t u v Nml_Nmlize by blast
moreover have "Nml (\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>)"
using 1 t u tu Nmlize_Src Nmlize_Trg Nml_HcompNml(1)
by presburger
moreover have "\<And>t. Arr t \<Longrightarrow> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>Dom t\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using t Nmlize_Vcomp_Arr_Dom by simp
moreover have "Dom \<^bold>\<lfloor>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom \<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<^bold>\<rfloor>"
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>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>"]
by auto
qed
thus ?thesis
using t u v Nmlize_Vcomp_Arr_Dom VcompNml_HcompNml Nml_Nmlize
by simp
qed
moreover have "\<lbrace>\<^bold>\<a>\<^bold>[t, u, v\<^bold>]\<rbrace> = \<lbrace>(t \<^bold>\<star> u \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<rbrace>"
using t u v tu uv Ide_Dom comp_cod_arr ide_eval_Ide \<alpha>_def
apply (simp add: eval_simps')
using assoc_is_natural_1 arr_eval_Arr eval_simps'(2-4) by presburger
ultimately show "coherent \<^bold>\<a>\<^bold>[t, u, v\<^bold>]" by argo
qed
show "Arr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \<Longrightarrow> coherent \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]"
proof -
assume tuv: "Arr \<^bold>\<a>\<^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>\<star> u) \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>])"
proof -
have "Arr ((t \<^bold>\<star> u) \<^bold>\<star> v) \<and> coherent ((t \<^bold>\<star> u) \<^bold>\<star> 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>\<a>\<^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>\<a>\<^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>\<star> u) \<^bold>\<star> v) = Cod \<^bold>\<a>\<^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>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] (((t \<^bold>\<star> u) \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>])"
using t u v tu uv Ide_Dom by simp
moreover have "\<^bold>\<lfloor>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<^bold>\<rfloor> = \<^bold>\<lfloor>((t \<^bold>\<star> u) \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<^bold>\<rfloor>"
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>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>"]
using Nmlize.simps(3-4,10) by presburger
moreover have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = \<lbrace>((t \<^bold>\<star> u) \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<rbrace>"
proof -
have 1: "VVV.arr (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
using tuv \<alpha>'.preserves_reflects_arr arr_eval_Arr eval.simps(10)
by (metis (no_types, lifting))
have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<rbrace> = ((\<lbrace>t\<rbrace> \<star> \<lbrace>u\<rbrace>) \<star> \<lbrace>v\<rbrace>) \<cdot> \<a>\<^sup>-\<^sup>1[dom \<lbrace>t\<rbrace>, dom \<lbrace>u\<rbrace>, dom \<lbrace>v\<rbrace>]"
proof -
have "VVV.arr (\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"
using tuv \<alpha>'.preserves_reflects_arr arr_eval_Arr eval.simps(10)
by (metis (no_types, lifting))
thus ?thesis
using t u v \<alpha>'.is_natural_1 [of "(\<lbrace>t\<rbrace>, \<lbrace>u\<rbrace>, \<lbrace>v\<rbrace>)"] HoHV_def \<a>'_def
VVV.dom_simp
by simp
qed
also have "... = \<lbrace>((t \<^bold>\<star> u) \<^bold>\<star> v) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<rbrace>"
by (simp add: eval_simps'(4) t u v \<a>'_def)
finally show ?thesis by blast
qed
ultimately show "coherent \<^bold>\<a>\<^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>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
shows "\<lbrace>t\<rbrace> = \<lbrace>u\<rbrace>"
using assms coherence canonical_factorization by simp
text \<open>
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.
\<close>
lemma canonically_isomorphicI:
assumes "f = \<lbrace>t\<rbrace>" and "g = \<lbrace>u\<rbrace>" and "Ide t" and "Ide u" and "\<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
shows "f \<cong> g"
proof -
have "f \<cong> \<lbrace>t\<rbrace>"
using assms isomorphic_reflexive ide_eval_Ide by blast
also have "... \<cong> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>"
proof -
have "\<guillemotleft>\<lbrace>t\<^bold>\<down>\<rbrace> : \<lbrace>t\<rbrace> \<Rightarrow> \<lbrace>\<^bold>\<lfloor>t\<^bold>\<rfloor>\<rbrace>\<guillemotright> \<and> iso \<lbrace>t\<^bold>\<down>\<rbrace>"
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 "... \<cong> \<lbrace>\<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace>"
using assms isomorphic_reflexive
by (simp add: Ide_Nmlize_Ide ide_eval_Ide)
also have "... \<cong> \<lbrace>u\<rbrace>"
proof -
have "\<guillemotleft>\<lbrace>u\<^bold>\<down>\<rbrace> : \<lbrace>u\<rbrace> \<Rightarrow> \<lbrace>\<^bold>\<lfloor>u\<^bold>\<rfloor>\<rbrace>\<guillemotright> \<and> iso \<lbrace>u\<^bold>\<down>\<rbrace>"
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 "... \<cong> g"
using assms isomorphic_reflexive ide_eval_Ide by blast
finally show ?thesis by simp
qed
end
end
diff --git a/thys/Bicategory/ConcreteBicategory.thy b/thys/Bicategory/ConcreteBicategory.thy
--- a/thys/Bicategory/ConcreteBicategory.thy
+++ b/thys/Bicategory/ConcreteBicategory.thy
@@ -1,1732 +1,1732 @@
(* Title: ConcreteBicategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2021
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Concrete Bicategories"
text \<open>
The locale \<open>concrete_bicategory\<close> defined in this section provides a uniform way to construct
a bicategory from extrinsically specified data comprising: a set of \<open>Obj\<close> of ``objects'',
a ``hom-category'' \<open>Hom A B\<close> for each pair of objects \<open>A\<close> and \<open>B\<close>, an ``identity arrow''
\<open>Id A \<in> Hom A A\<close> for each object \<open>A\<close>, ``horizontal composition'' functors
\<open>Comp C B A : Hom B C \<times> Hom A B \<rightarrow> Hom A C\<close> indexed by triples of objects,
together with unit and associativity isomorphisms; the latter subject to naturality and
coherence conditions.
We show that the bicategory produced by the construction relates to the given data in the
expected fashion: the objects of the bicategory are in bijective correspondence with the
given set \<open>Obj\<close>, the hom-categories of the bicategory are isomorphic to the given
categories \<open>Hom A B\<close>, the horizontal composition of the bicategory agrees with the given
compositions \<open>Comp C B A\<close>, and the unit and associativity 2-cells of the bicategory are
directly defined in terms of the given unit and associativity isomorphisms.
\<close>
theory ConcreteBicategory
imports Bicategory.Bicategory
begin
locale concrete_bicategory =
fixes Obj :: "'o set"
and Hom :: "'o \<Rightarrow> 'o \<Rightarrow> 'a comp"
and Id :: "'o \<Rightarrow> 'a"
and Comp :: "'o \<Rightarrow> 'o \<Rightarrow> 'o \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow>'a"
and Unit :: "'o \<Rightarrow> 'a"
and Assoc :: "'o \<Rightarrow> 'o \<Rightarrow> 'o \<Rightarrow> 'o \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
assumes category_Hom: "\<lbrakk> A \<in> Obj; B \<in> Obj \<rbrakk> \<Longrightarrow> category (Hom A B)"
and binary_functor_Comp:
"\<lbrakk> A \<in> Obj; B \<in> Obj; C \<in> Obj \<rbrakk>
\<Longrightarrow> binary_functor (Hom B C) (Hom A B) (Hom A C) (\<lambda>(f, g). Comp C B A f g)"
and ide_Id: "A \<in> Obj \<Longrightarrow> partial_magma.ide (Hom A A) (Id A)"
and Unit_in_hom:
"A \<in> Obj \<Longrightarrow>
partial_magma.in_hom (Hom A A) (Unit A) (Comp A A A (Id A) (Id A)) (Id A)"
and iso_Unit: "A \<in> Obj \<Longrightarrow> category.iso (Hom A A) (Unit A)"
and natural_isomorphism_Assoc:
"\<lbrakk> A \<in> Obj; B \<in> Obj; C \<in> Obj; D \<in> Obj \<rbrakk>
\<Longrightarrow> natural_isomorphism
(product_category.comp
(Hom C D) (product_category.comp (Hom B C) (Hom A B))) (Hom A D)
(\<lambda>(f, g, h). Comp D B A (Comp D C B f g) h)
(\<lambda>(f, g, h). Comp D C A f (Comp C B A g h))
(\<lambda>(f, g, h). Assoc D C B A f g h)"
and left_unit_Id:
"\<And>A B. \<lbrakk> A \<in> Obj; B \<in> Obj \<rbrakk>
\<Longrightarrow> fully_faithful_functor (Hom A B) (Hom A B)
(\<lambda>f. if partial_magma.arr (Hom A B) f
then Comp B B A (Id B) f
else partial_magma.null (Hom A B))"
and right_unit_Id:
"\<And>A B. \<lbrakk> A \<in> Obj; B \<in> Obj \<rbrakk>
\<Longrightarrow> fully_faithful_functor (Hom A B) (Hom A B)
(\<lambda>f. if partial_magma.arr (Hom A B) f
then Comp B A A f (Id A)
else partial_magma.null (Hom A B))"
and pentagon:
"\<And>A B C D E f g h k.
\<lbrakk> A \<in> Obj; B \<in> Obj; C \<in> Obj; D \<in> Obj; E \<in> Obj;
partial_magma.ide (Hom D E) f; partial_magma.ide (Hom C D) g;
partial_magma.ide (Hom B C) h; partial_magma.ide (Hom A B) k \<rbrakk> \<Longrightarrow>
Hom A E (Comp E D A f (Assoc D C B A g h k))
(Hom A E (Assoc E D B A f (Comp D C B g h) k)
(Comp E B A (Assoc E D C B f g h) k)) =
Hom A E (Assoc E D C A f g (Comp C B A h k))
(Assoc E C B A (Comp E D C f g) h k)"
begin
text \<open>
We first construct the vertical category.
Arrows are terms of the form \<open>MkCell A B \<mu>\<close>, where \<open>A \<in> Obj\<close>, \<open>B \<in> Obj\<close>, and where \<open>\<mu>\<close>
is an arrow of \<open>Hom A B\<close>.
Composition requires agreement of the ``source'' \<open>A\<close> and ``target'' \<open>B\<close> components,
and is then defined in terms of composition within \<open>Hom A B\<close>.
\<close>
datatype ('oo, 'aa) cell =
Null
| MkCell 'oo 'oo 'aa
abbreviation MkObj :: "'o \<Rightarrow> ('o, 'a) cell"
where "MkObj A \<equiv> MkCell A A (Id A)"
fun Src :: "('o, 'a) cell \<Rightarrow> 'o"
where "Src (MkCell A _ _) = A"
| "Src _ = undefined"
fun Trg
where "Trg (MkCell _ B _) = B"
| "Trg _ = undefined"
fun Map
where "Map (MkCell _ _ F) = F"
| "Map _ = undefined"
abbreviation Cell
where "Cell \<mu> \<equiv> \<mu> \<noteq> Null \<and> Src \<mu> \<in> Obj \<and> Trg \<mu> \<in> Obj \<and>
partial_magma.arr (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
definition vcomp
where "vcomp \<mu> \<nu> \<equiv> if Cell \<mu> \<and> Cell \<nu> \<and> Src \<mu> = Src \<nu> \<and> Trg \<mu> = Trg \<nu> \<and>
partial_magma.seq (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>) (Map \<nu>)
then MkCell (Src \<mu>) (Trg \<mu>) (Hom (Src \<mu>) (Trg \<mu>) (Map \<mu>) (Map \<nu>))
else Null"
interpretation partial_magma vcomp
by (metis partial_magma_def vcomp_def)
lemma null_char:
shows "null = Null"
using comp_null(1) vcomp_def by metis
lemma MkCell_Map:
assumes "\<mu> \<noteq> null"
shows "\<mu> = MkCell (Src \<mu>) (Trg \<mu>) (Map \<mu>)"
using assms null_char
by (metis Map.simps(1) Src.elims Trg.simps(1))
(*
* We need to name the following fact so that it does not get
* hidden when we interpret the category locale, because it is still
* used in some subsequent proofs.
*)
lemma ide_char'':
shows "ide \<mu> \<longleftrightarrow> Cell \<mu> \<and> partial_magma.ide (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
proof
show "ide \<mu> \<Longrightarrow> Cell \<mu> \<and> partial_magma.ide (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
proof
assume \<mu>: "ide \<mu>"
show 1: "Cell \<mu>"
by (metis \<mu> ide_def vcomp_def)
interpret Hom: category "Hom (Src \<mu>) (Trg \<mu>)"
using 1 category_Hom by simp
let ?\<nu> = "MkCell (Src \<mu>) (Trg \<mu>) (Hom.dom (Map \<mu>))"
have "vcomp \<mu> ?\<nu> = MkCell (Src \<mu>) (Trg \<mu>) (Map \<mu>)"
using 1 vcomp_def Hom.comp_arr_dom by simp
moreover have "vcomp \<mu> ?\<nu> = ?\<nu>"
using \<mu> ide_def null_char
by (metis MkCell_Map calculation)
ultimately show "Hom.ide (Map \<mu>)"
using 1 Hom.ide_dom by fastforce
qed
show "Cell \<mu> \<and> partial_magma.ide (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>) \<Longrightarrow> ide \<mu>"
proof -
assume \<mu>: "Cell \<mu> \<and> partial_magma.ide (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
interpret Hom: category "Hom (Src \<mu>) (Trg \<mu>)"
using \<mu> category_Hom by simp
show "ide \<mu>"
proof -
have "vcomp \<mu> \<mu> \<noteq> null"
using \<mu> vcomp_def null_char by simp
moreover have "\<And>\<nu>. vcomp \<nu> \<mu> \<noteq> null \<Longrightarrow> vcomp \<nu> \<mu> = \<nu>"
by (metis (full_types) Hom.comp_arr_ide \<mu> MkCell_Map vcomp_def null_char)
moreover have "\<And>\<nu>. vcomp \<mu> \<nu> \<noteq> null \<Longrightarrow> vcomp \<mu> \<nu> = \<nu>"
by (metis Hom.comp_ide_arr MkCell_Map \<mu> null_char vcomp_def)
ultimately show ?thesis
unfolding ide_def by simp
qed
qed
qed
lemma MkCell_in_domains:
assumes "Cell \<mu>"
shows "MkCell (Src \<mu>) (Trg \<mu>) (partial_magma.dom (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>))
\<in> domains \<mu>"
proof -
interpret Hom: category "Hom (Src \<mu>) (Trg \<mu>)"
using assms category_Hom by simp
let ?\<nu> = "MkCell (Src \<mu>) (Trg \<mu>) (Hom.dom (Map \<mu>))"
have "ide ?\<nu>"
using assms ide_char'' Hom.arr_dom Hom.ide_dom by simp
moreover have "vcomp \<mu> ?\<nu> = \<mu>"
unfolding vcomp_def
using assms Hom.comp_arr_dom MkCell_Map null_char by auto
ultimately show ?thesis
using domains_def by (simp add: assms null_char)
qed
lemma MkCell_in_codomains:
assumes "Cell \<mu>"
shows "MkCell (Src \<mu>) (Trg \<mu>) (partial_magma.cod (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>))
\<in> codomains \<mu>"
proof -
interpret Hom: category "Hom (Src \<mu>) (Trg \<mu>)"
using assms category_Hom by simp
let ?\<nu> = "MkCell (Src \<mu>) (Trg \<mu>) (Hom.cod (Map \<mu>))"
have "ide ?\<nu>"
using assms ide_char'' Hom.arr_dom Hom.ide_dom by simp
moreover have "vcomp ?\<nu> \<mu> = \<mu>"
unfolding vcomp_def
using assms Hom.comp_cod_arr MkCell_Map null_char by auto
ultimately show ?thesis
using codomains_def by (simp add: assms null_char)
qed
lemma has_domain_char:
shows "domains \<mu> \<noteq> {} \<longleftrightarrow> Cell \<mu>"
proof -
have "\<not> Cell \<mu> \<Longrightarrow> domains \<mu> = {}"
using vcomp_def domains_def null_char by auto
moreover have
"Cell \<mu> \<Longrightarrow> MkCell (Src \<mu>) (Trg \<mu>) (partial_magma.dom (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>))
\<in> domains \<mu>"
using MkCell_in_domains by simp
ultimately show ?thesis by auto
qed
lemma has_codomain_char:
shows "codomains \<mu> \<noteq> {} \<longleftrightarrow> Cell \<mu>"
proof -
have "\<not> Cell \<mu> \<Longrightarrow> codomains \<mu> = {}"
using vcomp_def codomains_def null_char by auto
moreover have
"Cell \<mu> \<Longrightarrow> MkCell (Src \<mu>) (Trg \<mu>) (partial_magma.cod (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>))
\<in> codomains \<mu>"
using MkCell_in_codomains by simp
ultimately show ?thesis by auto
qed
lemma arr_char:
shows "arr \<mu> \<longleftrightarrow> Cell \<mu>"
using arr_def has_domain_char has_codomain_char by simp
lemma ide_char''':
shows "ide \<mu> \<longleftrightarrow> arr \<mu> \<and> partial_magma.ide (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
using ide_char'' arr_char by simp
lemma seq_char:
shows "seq \<mu> \<nu> \<longleftrightarrow> Cell \<mu> \<and> Cell \<nu> \<and> Src \<mu> = Src \<nu> \<and> Trg \<mu> = Trg \<nu> \<and>
partial_magma.seq (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>) (Map \<nu>)"
using arr_char vcomp_def by auto
lemma vcomp_char:
shows "vcomp \<mu> \<nu> = (if seq \<mu> \<nu> then
MkCell (Src \<mu>) (Trg \<mu>) (Hom (Src \<mu>) (Trg \<mu>) (Map \<mu>) (Map \<nu>))
else null)"
by (metis null_char seq_char vcomp_def)
interpretation category vcomp
proof
show "\<And>g f. vcomp g f \<noteq> null \<Longrightarrow> seq g f"
using seq_char null_char vcomp_def by metis
show "\<And>f. (domains f \<noteq> {}) = (codomains f \<noteq> {})"
using has_domain_char has_codomain_char by simp
show "\<And>h g f. \<lbrakk>seq h g; seq (vcomp h g) f\<rbrakk> \<Longrightarrow> seq g f"
using vcomp_def
apply (unfold seq_char, intro conjI)
apply auto
by (meson category.match_1 category_Hom)
show "\<And>h g f. \<lbrakk>seq h (vcomp g f); seq g f\<rbrakk> \<Longrightarrow> seq h g"
using vcomp_def
apply (unfold seq_char, intro conjI)
apply auto
by (meson category.match_2 category_Hom)
show "\<And>g f h. \<lbrakk>seq g f; seq h g\<rbrakk> \<Longrightarrow> seq (vcomp h g) f"
using vcomp_def
apply (unfold seq_char, intro conjI)
apply auto
by (meson category.match_3 category_Hom)
show "\<And>g f h. \<lbrakk>seq g f; seq h g\<rbrakk> \<Longrightarrow> vcomp (vcomp h g) f = vcomp h (vcomp g f)"
proof -
fix f g h
assume fg: "seq g f" and gh: "seq h g"
interpret Hom: category \<open>Hom (Src f) (Trg f)\<close>
using fg seq_char category_Hom by simp
have "vcomp (vcomp h g) f =
MkCell (Src f) (Trg f)
(Hom (Src f) (Trg f)
(Hom (Src f) (Trg f) (Map h) (Map g)) (Map f))"
using fg gh vcomp_char seq_char null_char Hom.match_3 by auto
also have "... =
MkCell (Src f) (Trg f)
(Hom (Src f) (Trg f) (Map h)
(Hom (Src f) (Trg f) (Map g) (Map f)))"
using fg gh seq_char Hom.comp_assoc by simp
also have "... = vcomp h (vcomp g f)"
using fg gh vcomp_char seq_char null_char Hom.match_4 by auto
finally show "vcomp (vcomp h g) f = vcomp h (vcomp g f)"
by blast
qed
qed
lemma arr_eqI:
assumes "arr f" and "arr f'"
and "Src f = Src f'" and "Trg f = Trg f'" and "Map f = Map f'"
shows "f = f'"
using arr_char MkCell_Map assms null_char by metis
lemma dom_char:
shows "dom \<mu> = (if arr \<mu> then
MkCell (Src \<mu>) (Trg \<mu>) (partial_magma.dom (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>))
else Null)"
by (metis MkCell_in_domains arr_char dom_in_domains domain_unique has_domain_iff_arr
dom_def null_char)
lemma cod_char:
shows "cod \<mu> = (if arr \<mu> then
MkCell (Src \<mu>) (Trg \<mu>) (partial_magma.cod (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>))
else Null)"
by (metis MkCell_in_codomains arr_char cod_def cod_in_codomains codomain_unique
has_codomain_iff_arr null_char)
lemma Src_vcomp [simp]:
assumes "seq \<mu> \<nu>"
shows "Src (vcomp \<mu> \<nu>) = Src \<mu>"
using assms seq_char vcomp_def by auto
lemma Trg_vcomp [simp]:
assumes "seq \<mu> \<nu>"
shows "Trg (vcomp \<mu> \<nu>) = Trg \<mu>"
using assms seq_char vcomp_def by auto
lemma Map_vcomp [simp]:
assumes "seq \<mu> \<nu>"
shows "Map (vcomp \<mu> \<nu>) = Hom (Src \<mu>) (Trg \<mu>) (Map \<mu>) (Map \<nu>)"
using assms seq_char vcomp_def by auto
lemma arr_MkCell [simp]:
assumes "A \<in> Obj" and "B \<in> Obj" and "partial_magma.arr (Hom A B) f"
shows "arr (MkCell A B f)"
using assms arr_char by simp
lemma dom_MkCell [simp]:
assumes "arr (MkCell A B f)"
shows "dom (MkCell A B f) = MkCell A B (partial_magma.dom (Hom A B) f)"
using assms arr_char dom_char by simp
lemma cod_MkCell [simp]:
assumes "arr (MkCell A B f)"
shows "cod (MkCell A B f) = MkCell A B (partial_magma.cod (Hom A B) f)"
using assms arr_char cod_char by simp
lemma iso_char:
shows "iso \<mu> \<longleftrightarrow> arr \<mu> \<and> category.iso (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
proof
assume \<mu>: "iso \<mu>"
have 1: "arr \<mu>" using \<mu> by blast
interpret Hom: category \<open>Hom (Src \<mu>) (Trg \<mu>)\<close>
using 1 arr_char category_Hom by simp
have 2: "Hom.iso (Map \<mu>)"
proof -
obtain \<nu> where \<nu>: "inverse_arrows \<mu> \<nu>"
using \<mu> by blast
have "Hom.inverse_arrows (Map \<mu>) (Map \<nu>)"
proof
show "Hom.ide (Hom (Src \<mu>) (Trg \<mu>) (Map \<mu>) (Map \<nu>))"
using \<nu> ide_char'' Src_vcomp Trg_vcomp ideD(1) vcomp_char Map_vcomp
by (metis inverse_arrowsE)
show "Hom.ide (Hom (Src \<mu>) (Trg \<mu>) (Map \<nu>) (Map \<mu>))"
proof -
have 1: "ide (vcomp \<nu> \<mu>)"
using \<nu> by auto
hence "Hom.ide (Map (vcomp \<nu> \<mu>))"
using ide_char'' Src_vcomp Trg_vcomp ideD(1) seq_char by metis
thus "Hom.ide (Hom (Src \<mu>) (Trg \<mu>) (Map \<nu>) (Map \<mu>))"
using \<nu> 1 vcomp_char Map.simps(1) seq_char ideD(1)
by (metis (no_types, lifting))
qed
qed
thus ?thesis by auto
qed
show "arr \<mu> \<and> Hom.iso (Map \<mu>)"
using 1 2 by simp
next
assume \<mu>: "arr \<mu> \<and> category.iso (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
interpret Hom: category \<open>Hom (Src \<mu>) (Trg \<mu>)\<close>
using \<mu> arr_char category_Hom by simp
obtain f where f: "Hom.inverse_arrows (Map \<mu>) f"
using \<mu> by auto
let ?\<nu> = "MkCell (Src \<mu>) (Trg \<mu>) f"
have 1: "arr ?\<nu>"
using \<mu> f arr_char by auto
have "inverse_arrows \<mu> (MkCell (Src \<mu>) (Trg \<mu>) f)"
using \<mu> f 1 arr_char ide_char'' vcomp_def
apply (intro inverse_arrowsI) by auto
thus "iso \<mu>" by auto
qed
text \<open>
Next, we equip each arrow with a source and a target, and show that these assignments
are functorial.
\<close>
definition src
where "src \<mu> \<equiv> if arr \<mu> then MkObj (Src \<mu>) else null"
definition trg
where "trg \<mu> \<equiv> if arr \<mu> then MkObj (Trg \<mu>) else null"
lemma src_MkCell [simp]:
assumes "arr (MkCell A B f)"
shows "src (MkCell A B f) = MkObj A"
using assms src_def by simp
lemma trg_MkCell [simp]:
assumes "arr (MkCell A B f)"
shows "trg (MkCell A B f) = MkObj B"
using assms trg_def by simp
lemma src_dom:
assumes "arr \<mu>"
shows "src (dom \<mu>) = src \<mu>"
using assms dom_char src_def arr_char arr_dom by auto
lemma src_cod:
assumes "arr \<mu>"
shows "src (cod \<mu>) = src \<mu>"
using assms cod_char src_def arr_char arr_cod by auto
lemma trg_dom:
assumes "arr \<mu>"
shows "trg (dom \<mu>) = trg \<mu>"
using assms dom_char trg_def arr_char arr_dom by auto
lemma trg_cod:
assumes "arr \<mu>"
shows "trg (cod \<mu>) = trg \<mu>"
using assms cod_char trg_def arr_char arr_cod by auto
lemma Src_src [simp]:
assumes "arr \<mu>"
shows "Src (src \<mu>) = Src \<mu>"
using assms src_def by simp
lemma Trg_src [simp]:
assumes "arr \<mu>"
shows "Trg (src \<mu>) = Src \<mu>"
using assms src_def by simp
lemma Map_src [simp]:
assumes "arr \<mu>"
shows "Map (src \<mu>) = Id (Src \<mu>)"
using assms src_def by simp
lemma Src_trg [simp]:
assumes "arr \<mu>"
shows "Src (trg \<mu>) = Trg \<mu>"
using assms trg_def by simp
lemma Trg_trg [simp]:
assumes "arr \<mu>"
shows "Trg (trg \<mu>) = Trg \<mu>"
using assms trg_def by simp
lemma Map_trg [simp]:
assumes "arr \<mu>"
shows "Map (trg \<mu>) = Id (Trg \<mu>)"
using assms trg_def by simp
lemma Src_dom [simp]:
assumes "arr \<mu>"
shows "Src (dom \<mu>) = Src \<mu>"
using assms dom_char src_def arr_char arr_dom by auto
lemma Src_cod [simp]:
assumes "arr \<mu>"
shows "Src (cod \<mu>) = Src \<mu>"
using assms src_cod src_def arr_char arr_cod by auto
lemma Trg_dom [simp]:
assumes "arr \<mu>"
shows "Trg (dom \<mu>) = Trg \<mu>"
using assms dom_char trg_def arr_char arr_dom by auto
lemma Trg_cod [simp]:
assumes "arr \<mu>"
shows "Trg (cod \<mu>) = Trg \<mu>"
using assms cod_char trg_def arr_char arr_cod by auto
lemma Map_dom [simp]:
assumes "arr \<mu>"
shows "Map (dom \<mu>) = partial_magma.dom (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
using assms by (simp add: dom_char)
lemma Map_cod [simp]:
assumes "arr \<mu>"
shows "Map (cod \<mu>) = partial_magma.cod (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
using assms by (simp add: cod_char)
lemma ide_MkObj:
assumes "A \<in> Obj"
shows "ide (MkObj A)"
using assms ide_char'
by (metis Map.simps(1) Src.simps(1) Trg.simps(1) category.ideD(1) cell.simps(2)
category_Hom ide_char'' ide_Id)
interpretation src: "functor" vcomp vcomp src
using src_def arr_char Map.simps(1) Src.simps(1) Trg.simps(1) arr_dom category.ideD(1)
ide_MkObj src_dom src_cod ide_Id
apply unfold_locales
apply auto[1]
apply (simp add: category.ideD(1) category_Hom)
apply auto[2]
proof -
fix g :: "('o, 'a) cell" and f :: "('o, 'a) cell"
assume fg: "seq g f"
thus "src (vcomp g f) = vcomp (src g) (src f)"
using arr_char ide_MkObj src_dom src_cod src_def
by (metis Src_vcomp comp_ide_self seqE)
qed
interpretation trg: "functor" vcomp vcomp trg
using trg_def arr_char Map.simps(1) Src.simps(1) Trg.simps(1) arr_dom category.ideD(1)
ide_MkObj trg_dom trg_cod ide_Id
apply unfold_locales
apply auto[1]
apply (simp add: category.ideD(1) category_Hom)
apply auto[2]
proof -
fix g :: "('o, 'a) cell" and f :: "('o, 'a) cell"
assume fg: "seq g f"
thus "trg (vcomp g f) = vcomp (trg g) (trg f)"
using arr_char ide_MkObj trg_dom trg_cod trg_def
by (metis Trg_vcomp comp_ide_self seqE)
qed
interpretation H: horizontal_homs vcomp src trg
using ide_MkObj arr_char src_def trg_def src.preserves_arr trg.preserves_arr
by unfold_locales auto
lemma obj_MkObj:
assumes "A \<in> Obj"
shows "H.obj (MkObj A)"
using assms src_def H.obj_def ide_MkObj by simp
lemma MkCell_in_hom [intro]:
assumes "A \<in> Obj" and "B \<in> Obj" and "partial_magma.arr (Hom A B) f"
shows "H.in_hhom (MkCell A B f) (MkObj A) (MkObj B)"
and "\<guillemotleft>MkCell A B f : MkCell A B (partial_magma.dom (Hom A B) f)
\<Rightarrow> MkCell A B (partial_magma.cod (Hom A B) f)\<guillemotright>"
using assms by auto
text \<open>
Horizontal composition of horizontally composable arrows is now defined by applying
the given function \<open>Comp\<close> to the ``Map'' components.
\<close>
definition hcomp
where "hcomp \<mu> \<nu> \<equiv> if arr \<mu> \<and> arr \<nu> \<and> src \<mu> = trg \<nu> then
MkCell (Src \<nu>) (Trg \<mu>) (Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>))
else
null"
lemma arr_hcomp:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "arr (hcomp \<mu> \<nu>)"
and "dom (hcomp \<mu> \<nu>) = hcomp (dom \<mu>) (dom \<nu>)"
and "cod (hcomp \<mu> \<nu>) = hcomp (cod \<mu>) (cod \<nu>)"
proof -
have 1: "hcomp \<mu> \<nu> =
MkCell (Src \<nu>) (Trg \<mu>) (Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>))"
using assms hcomp_def by simp
have 2: "Src \<mu> = Trg \<nu>"
using assms src_def trg_def by simp
interpret Hom_\<mu>: category \<open>Hom (Src \<mu>) (Trg \<mu>)\<close>
using assms arr_char category_Hom by simp
interpret Hom_\<nu>: category \<open>Hom (Src \<nu>) (Trg \<nu>)\<close>
using assms arr_char category_Hom by simp
interpret Hom_\<mu>\<nu>: category \<open>Hom (Src \<nu>) (Trg \<mu>)\<close>
using assms arr_char category_Hom by simp
interpret Comp: binary_functor
\<open>Hom (Trg \<nu>) (Trg \<mu>)\<close> \<open>Hom (Src \<nu>) (Trg \<nu>)\<close> \<open>Hom (Src \<nu>) (Trg \<mu>)\<close>
\<open>\<lambda>(f, g). Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) f g\<close>
using assms arr_char 2 binary_functor_Comp [of "Src \<nu>" "Trg \<nu>" "Trg \<mu>"]
by simp
have 4: "Comp.A1xA2.arr (Map \<mu>, Map \<nu>)"
using assms 2 arr_char Comp.A1xA2.arr_char by simp
show 3: "arr (hcomp \<mu> \<nu>)"
using assms 1 2 4 arr_char Comp.preserves_arr [of "(Map \<mu>, Map \<nu>)"]
by simp
show "dom (hcomp \<mu> \<nu>) = hcomp (dom \<mu>) (dom \<nu>)"
proof -
have "dom (hcomp \<mu> \<nu>) =
MkCell (Src \<nu>) (Trg \<mu>)
(Hom_\<mu>\<nu>.dom (Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>)))"
using 1 3 dom_char by simp
moreover have "Hom_\<mu>\<nu>.dom (Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>)) =
Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Hom_\<mu>.dom (Map \<mu>)) (Hom_\<nu>.dom (Map \<nu>))"
using 2 Comp.preserves_dom \<open>Comp.A1xA2.arr (Map \<mu>, Map \<nu>)\<close> by force
ultimately show ?thesis
using assms 2 dom_char hcomp_def arr_dom
by auto metis
qed
show "cod (hcomp \<mu> \<nu>) = hcomp (cod \<mu>) (cod \<nu>)"
proof -
have "cod (hcomp \<mu> \<nu>) =
MkCell (Src \<nu>) (Trg \<mu>)
(Hom_\<mu>\<nu>.cod (Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>)))"
using 1 3 cod_char by simp
moreover have "Hom_\<mu>\<nu>.cod (Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>)) =
Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Hom_\<mu>.cod (Map \<mu>)) (Hom_\<nu>.cod (Map \<nu>))"
using 2 Comp.preserves_cod \<open>Comp.A1xA2.arr (Map \<mu>, Map \<nu>)\<close> by force
ultimately show "cod (hcomp \<mu> \<nu>) = hcomp (cod \<mu>) (cod \<nu>)"
using assms 2 cod_char hcomp_def arr_cod
by auto metis
qed
qed
lemma src_hcomp:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "src (hcomp \<mu> \<nu>) = src \<nu>"
using assms hcomp_def src_def arr_hcomp(1) by auto
lemma trg_hcomp:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "trg (hcomp \<mu> \<nu>) = trg \<mu>"
using assms hcomp_def trg_def arr_hcomp(1) by auto
lemma Src_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "Src (hcomp \<mu> \<nu>) = Src \<nu>"
using assms hcomp_def by simp
lemma Trg_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "Trg (hcomp \<mu> \<nu>) = Trg \<mu>"
using assms hcomp_def by simp
lemma Map_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "Map (hcomp \<mu> \<nu>) = Comp (Trg \<mu>) (Trg \<nu>) (Src \<nu>) (Map \<mu>) (Map \<nu>)"
using assms hcomp_def by simp
lemma hcomp_vcomp:
assumes "H.VV.seq g f"
shows "hcomp (fst (H.VV.comp g f)) (snd (H.VV.comp g f)) =
vcomp (hcomp (fst g) (snd g)) (hcomp (fst f) (snd f))"
proof -
let ?f1 = "fst f" and ?f2 = "snd f" and ?g1 = "fst g" and ?g2 = "snd g"
have 1: "Src ?f1 \<in> Obj \<and> Trg ?f1 \<in> Obj \<and> Src ?f2 \<in> Obj \<and> Trg ?f2 \<in> Obj \<and>
Src ?g1 \<in> Obj \<and> Trg ?g1 \<in> Obj \<and> Src ?g2 \<in> Obj \<and> Trg ?g2 \<in> Obj"
using assms arr_char H.VV.arrE by blast
interpret Hom_f1: category \<open>Hom (Src ?f1) (Trg ?f1)\<close>
using assms 1 category_Hom by simp
interpret Hom_f2: category \<open>Hom (Src ?f2) (Trg ?f2)\<close>
using assms 1 category_Hom by simp
interpret Hom_g1: category \<open>Hom (Src ?g1) (Trg ?g1)\<close>
using assms 1 category_Hom by simp
interpret Hom_g2: category \<open>Hom (Src ?g2) (Trg ?g2)\<close>
using assms 1 category_Hom by simp
interpret Hom_f: category \<open>Hom (Src ?f2) (Trg ?f1)\<close>
using assms 1 category_Hom by simp
interpret Hom_g: category \<open>Hom (Src ?g2) (Trg ?g1)\<close>
using assms 1 category_Hom by simp
interpret Comp_f: binary_functor \<open>Hom (Trg ?f2) (Trg ?f1)\<close> \<open>Hom (Src ?f2) (Trg ?f2)\<close>
\<open>Hom (Src ?f2) (Trg ?f1)\<close>
\<open>\<lambda>(fa, g). Comp (Trg ?f1) (Trg ?f2) (Src ?f2) fa g\<close>
using assms 1 arr_char binary_functor_Comp by simp
have "hcomp (fst (H.VV.comp g f)) (snd (H.VV.comp g f)) =
MkCell (Src (snd (H.VV.comp g f))) (Trg (fst (H.VV.comp g f)))
(Comp (Trg (fst (H.VV.comp g f))) (Trg (snd (H.VV.comp g f)))
(Src (snd (H.VV.comp g f)))
(Map (fst (H.VV.comp g f))) (Map (snd (H.VV.comp g f))))"
using assms hcomp_def H.VV.arrE
by (metis (no_types, lifting))
also have "... = MkCell (Src ?f2) (Trg ?f1)
(Hom (Src ?f2) (Trg ?f1)
(Comp (Trg ?f1) (Trg ?g2) (Src ?f2) (Map ?g1) (Map ?g2))
(Comp (Trg ?f1) (Trg ?f2) (Src ?f2) (Map ?f1) (Map ?f2)))"
proof -
have "Src (snd (H.VV.comp g f)) = Src ?f2"
- using assms arr_char src_def H.VV.comp_char H.VV.seq_char
+ using assms arr_char src_def H.VV.comp_char H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) H.vseq_implies_hpar(1) Src.simps(1) H.VV.arrE
H.VV.inclusion H.VxV.comp_arr_dom H.VxV.dom_comp H.VxV.seqE)
moreover have "Trg (fst (H.VV.comp g f)) = Trg ?f1"
- by (metis (no_types, lifting) H.VV.comp_arr_dom H.VV.comp_simp H.VV.seq_char
+ by (metis (no_types, lifting) H.VV.comp_arr_dom H.VV.comp_simp H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C
H.VxV.arr_char H.VxV.cod_comp H.VxV.comp_cod_arr H.VxV.seqE
H.vseq_implies_hpar(2) Src_trg assms)
moreover have
"Comp (Trg (fst (H.VV.comp g f))) (Trg (snd (H.VV.comp g f)))
(Src (snd (H.VV.comp g f)))
(Map (fst (H.VV.comp g f))) (Map (snd (H.VV.comp g f))) =
Hom (Src ?f2) (Trg ?f1)
(Comp (Trg ?f1) (Trg ?g2) (Src ?f2) (Map ?g1) (Map ?g2))
(Comp (Trg ?f1) (Trg ?f2) (Src ?f2) (Map ?f1) (Map ?f2))"
proof -
have "Comp (Trg (fst (H.VV.comp g f))) (Trg (snd (H.VV.comp g f)))
(Src (snd (H.VV.comp g f)))
(Map (fst (H.VV.comp g f))) (Map (snd (H.VV.comp g f))) =
Comp (Trg ?g1) (Trg ?g2) (Src ?g2)
(Map (vcomp ?g1 ?f1)) (Map (vcomp ?g2 ?f2))"
- using assms H.VV.comp_char H.VV.arr_char H.VxV.comp_char by auto (* 10 sec *)
+ using assms H.VV.comp_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VxV.comp_char by auto (* 10 sec *)
also have "... = Comp (Trg ?g1) (Trg ?g2) (Src ?g2)
(Hom (Src ?g1) (Trg ?g1) (Map ?g1) (Map ?f1))
(Hom (Src ?g2) (Trg ?g2) (Map ?g2) (Map ?f2))"
- using assms H.VV.seq_char Map_vcomp H.VxV.seq_char by auto
+ using assms H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C Map_vcomp H.VxV.seq_char by auto
also have "... = Hom (Src ?f2) (Trg ?f1)
(Comp (Trg ?f1) (Trg ?g2) (Src ?f2)
(Map ?g1) (Map ?g2))
(Comp (Trg ?f1) (Trg ?f2) (Src ?f2)
(Map ?f1) (Map ?f2))"
proof -
have 2: "Src ?g1 = Trg ?g2"
- using assms H.VV.arr_char [of g] src_def [of "?g1"] trg_def [of "?g2"] by auto
+ using assms H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C [of g] src_def [of "?g1"] trg_def [of "?g2"] by auto
have "Comp (Trg ?f1) (Trg ?f2) (Src ?f2)
(Hom (Trg ?g2) (Trg ?g1) (Map ?g1) (Map ?f1))
(Hom (Src ?g2) (Trg ?g2) (Map ?g2) (Map ?f2)) =
Hom (Src ?f2) (Trg ?f1)
(Comp (Trg ?f1) (Trg ?f2) (Src ?f2)
(Map ?g1) (Map ?g2))
(Comp (Trg ?f1) (Trg ?f2) (Src ?f2)
(Map ?f1) (Map ?f2))"
proof -
have "Comp_f.A1xA2.seq (Map ?g1, Map ?g2) (Map ?f1, Map ?f2)"
- using assms 2 H.VV.seq_char
+ using assms 2 H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) Comp_f.A1xA2.seq_char H.VxV.seqE
fst_conv seq_char snd_conv)
moreover have
"Comp_f.A1xA2.comp (Map ?g1, Map ?g2) (Map ?f1, Map ?f2) =
(Hom (Src ?g1) (Trg ?g1) (Map ?g1) (Map ?f1),
Hom (Src ?g2) (Trg ?g2) (Map ?g2) (Map ?f2))"
- using assms 2 H.VV.seq_char H.VxV.seqE seq_char
+ using assms 2 H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C H.VxV.seqE seq_char
by (metis (no_types, lifting) Comp_f.A1xA2.comp_char fst_conv snd_conv)
ultimately show ?thesis
by (metis 2 Comp_f.as_nat_trans.preserves_comp_2 old.prod.case)
qed
- thus ?thesis using assms 2 H.VV.seq_char H.VxV.seqE seq_char
+ thus ?thesis using assms 2 H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C H.VxV.seqE seq_char
by (metis (no_types, lifting))
qed
finally show ?thesis by blast
qed
ultimately show ?thesis by blast
qed
also have "... = vcomp (hcomp ?g1 ?g2) (hcomp ?f1 ?f2)"
proof -
have 2: "Trg ?g1 = Trg ?f1 \<and> Src ?g2 = Src ?f2 \<and> Src ?f1 = Trg ?f2"
- using assms seq_char H.VV.seq_char H.VxV.seqE
- by (metis (no_types, lifting) H.VV.arr_char Src_src Src_trg)
+ using assms seq_char H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C H.VxV.seqE
+ by (metis (no_types, lifting) H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C Src_src Src_trg)
have "Hom_f.seq (Comp (Trg ?f1) (Trg ?g2) (Src ?f2)
(Map ?g1) (Map ?g2))
(Comp (Trg ?f1) (Trg ?f2) (Src ?f2)
(Map ?f1) (Map ?f2))"
- by (metis (no_types, lifting) 2 Comp_f.A1xA2.seqI Comp_f.preserves_seq H.VV.seq_char
+ by (metis (no_types, lifting) 2 Comp_f.A1xA2.seqI Comp_f.preserves_seq H.VV.seq_char\<^sub>S\<^sub>b\<^sub>C
H.VxV.seqE arr_char assms case_prod_conv vcomp_def)
moreover have "?f1 \<noteq> Null \<and> ?f2 \<noteq> Null \<and> src ?f1 = trg ?f2 \<and>
?g1 \<noteq> Null \<and> ?g2 \<noteq> Null \<and> src ?g1 = trg ?g2"
- using assms H.VV.arr_char arr_char assms by blast
+ using assms H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char assms by blast
moreover have "Hom_f1.arr (Map ?f1) \<and> Hom_f2.arr (Map ?f2)"
using assms seq_char H.VV.arrE H.VV.seqE arr_char by fast
moreover have "Hom_g1.arr (Map ?g1) \<and> Hom_g2.arr (Map ?g2)"
using assms seq_char H.VV.arrE H.VV.seqE arr_char by meson
ultimately show ?thesis
using 1 2 arr_char hcomp_def vcomp_def by auto
qed
finally show ?thesis by simp
qed
interpretation H: "functor" H.VV.comp vcomp \<open>\<lambda>\<mu>\<nu>. hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
- using hcomp_def arr_hcomp hcomp_vcomp H.VV.arr_char H.VV.dom_char H.VV.cod_char
+ using hcomp_def arr_hcomp hcomp_vcomp H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VV.dom_char\<^sub>S\<^sub>b\<^sub>C H.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by unfold_locales auto
interpretation H: horizontal_composition vcomp hcomp src trg
using src_hcomp trg_hcomp arr_char hcomp_def null_char
by unfold_locales auto
lemma Map_obj:
assumes "H.obj a"
shows "Map a = Id (Src a)" and "Map a = Id (Trg a)"
using assms H.obj_def Map_src Map_trg H.obj_simps(3) by metis+
lemma MkCell_simps:
assumes "A \<in> Obj" and "B \<in> Obj" and "partial_magma.arr (Hom A B) f"
shows "arr (MkCell A B f)"
and "src (MkCell A B f) = MkObj A" and "trg (MkCell A B f) = MkObj B"
and "dom (MkCell A B f) = MkCell A B (partial_magma.dom (Hom A B) f)"
and "cod (MkCell A B f) = MkCell A B (partial_magma.cod (Hom A B) f)"
using assms MkCell_in_hom by auto
text \<open>
Next, define the associativities and show that they are the components of a
natural isomorphism.
\<close>
definition assoc
where "assoc f g h \<equiv>
if H.VVV.ide (f, g, h) then
MkCell (Src h) (Trg f)
(Assoc (Trg f) (Trg g) (Trg h) (Src h) (Map f) (Map g) (Map h))
else null"
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 "\<guillemotleft>assoc f g h : hcomp (hcomp f g) h \<Rightarrow> hcomp f (hcomp g h)\<guillemotright>"
proof -
let ?A = "Src h" and ?B = "Trg h" and ?C = "Trg g" and ?D = "Trg f"
interpret Hom_f: category \<open>Hom ?C ?D\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_g: category \<open>Hom ?B ?C\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_h: category \<open>Hom ?A ?B\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_gh: product_category \<open>Hom ?B ?C\<close> \<open>Hom ?A ?B\<close> ..
interpret Hom_f_gh: product_category \<open>Hom ?C ?D\<close> Hom_gh.comp ..
interpret Comp_fg: binary_functor \<open>Hom ?C ?D\<close> \<open>Hom ?B ?C\<close> \<open>Hom ?B ?D\<close>
\<open>\<lambda>(fa, g). Comp ?D ?C ?B fa g\<close>
using assms arr_char ide_char'' binary_functor_Comp [of ?B ?C ?D] by simp
interpret \<alpha>: natural_isomorphism Hom_f_gh.comp \<open>Hom (Src h) (Trg f)\<close>
\<open>\<lambda>(fa, ga, ha). Comp ?D ?B ?A (Comp ?D ?C ?B fa ga) ha\<close>
\<open>\<lambda>(fa, ga, ha). Comp ?D ?C ?A fa (Comp ?C ?B ?A ga ha)\<close>
\<open>\<lambda>(fa, ga, ha). Assoc ?D ?C ?B ?A fa ga ha\<close>
using assms ide_char arr_char natural_isomorphism_Assoc [of ?A ?B ?C ?D]
by blast
show ?thesis
proof
have 1: "Src f = Trg g \<and> Src g = Trg h"
using assms src_def trg_def by simp
have 2: "Hom_f.ide (Map f) \<and> Hom_g.ide (Map g) \<and> Hom_h.ide (Map h)"
using assms 1 ide_char' [of f] arr_char [of f]
by (simp add: ide_char'')
show 3: "arr (assoc f g h)"
unfolding assoc_def
- using assms arr_char ide_char'' H.VV.arr_char H.VVV.arr_char H.VVV.ide_char
+ using assms arr_char ide_char'' H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
src_def trg_def \<alpha>.preserves_reflects_arr [of "(Map f, Map g, Map h)"]
Hom_f_gh.arr_char Hom_gh.arr_char
by simp
show "dom (assoc f g h) = hcomp (hcomp f g) h"
proof -
have "arr (MkCell ?B ?D (Comp ?D ?C ?B (Map f) (Map g)))"
by (metis assms(1-2,4) 1 Map_hcomp Src_hcomp Trg_hcomp arr_MkCell arr_char
arr_hcomp(1) ideD(1))
moreover have "MkObj ?B = trg h"
using assms ide_char'' arr_char MkCell_Map null_char trg_MkCell by metis
ultimately show ?thesis
unfolding hcomp_def
- using assms 1 2 3 arr_char ide_char'' assoc_def dom_MkCell H.VV.arr_char
- H.VVV.arr_char H.VVV.ide_char src_MkCell trg_MkCell \<alpha>.preserves_dom
+ using assms 1 2 3 arr_char ide_char'' assoc_def dom_MkCell H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C src_MkCell trg_MkCell \<alpha>.preserves_dom
by force
qed
show "cod (assoc f g h) = hcomp f (hcomp g h)"
proof -
have "trg g = MkObj ?C"
using assms ide_char'' arr_char MkCell_Map null_char trg_MkCell by metis
thus ?thesis
unfolding hcomp_def
using assms 2 3 \<alpha>.preserves_cod src_MkCell trg_MkCell H.hseqI' hcomp_def
- assms arr_char ide_char'' assoc_def cod_MkCell H.VV.arr_char
- H.VVV.arr_char H.VVV.ide_char
+ assms arr_char ide_char'' assoc_def cod_MkCell H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
by force
qed
qed
qed
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 (assoc f g h)"
and "dom (assoc f g h) = hcomp (hcomp f g) h"
and "cod (assoc f g h) = hcomp f (hcomp g h)"
using assms assoc_in_hom by auto
lemma assoc_simps' [simp]:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "src (assoc f g h) = src h"
and "trg (assoc f g h) = trg f"
and "Src (assoc f g h) = Src h"
and "Trg (assoc f g h) = Trg f"
and "Map (assoc f g h) = Assoc (Trg f) (Trg g) (Trg h) (Src h) (Map f) (Map g) (Map h)"
proof -
show "src (assoc f g h) = src h"
using assms src_hcomp src_dom src_def src_MkCell assoc_simps(1) assoc_def
by (metis (no_types, lifting) ideD(1) not_arr_null)
show "trg (assoc f g h) = trg f"
using assms trg_hcomp trg_hcomp trg_cod H.hseqI'
by (metis assoc_simps(1,3) ideD(1))
show "Src (assoc f g h) = Src h"
using assms Src_hcomp Src_dom
by (metis (no_types, lifting) Src.simps(1) arr_char assoc_def assoc_simps(1) null_char)
show "Trg (assoc f g h) = Trg f"
using assms Trg_hcomp Trg_dom
by (metis (no_types, lifting) Trg.simps(1) arr_char assoc_def assoc_simps(1) null_char)
show "Map (assoc f g h) = Assoc (Trg f) (Trg g) (Trg h) (Src h) (Map f) (Map g) (Map h)"
using assms Map_hcomp Map_dom
by (metis (no_types, lifting) Map.simps(1) arr_char assoc_def assoc_simps(1) null_char)
qed
lemma iso_assoc:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "iso (assoc f g h)"
proof -
let ?A = "Src h" and ?B = "Trg h" and ?C = "Trg g" and ?D = "Trg f"
interpret Hom_f: category \<open>Hom ?C ?D\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_g: category \<open>Hom ?B ?C\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_h: category \<open>Hom ?A ?B\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_fg: product_category \<open>Hom ?C ?D\<close> \<open>Hom ?B ?C\<close> ..
interpret Hom_gh: product_category \<open>Hom ?B ?C\<close> \<open>Hom ?A ?B\<close> ..
interpret Hom_f_gh: product_category \<open>Hom ?C ?D\<close> Hom_gh.comp ..
interpret \<alpha>: natural_isomorphism Hom_f_gh.comp \<open>Hom (Src h) (Trg f)\<close>
\<open>\<lambda>(fa, ga, ha). Comp ?D ?B ?A (Comp ?D ?C ?B fa ga) ha\<close>
\<open>\<lambda>(fa, ga, ha). Comp ?D ?C ?A fa (Comp ?C ?B ?A ga ha)\<close>
\<open>\<lambda>(fa, ga, ha). Assoc ?D ?C ?B ?A fa ga ha\<close>
using assms ide_char arr_char natural_isomorphism_Assoc [of ?A ?B ?C ?D]
by blast
show ?thesis
using assms \<alpha>.components_are_iso [of "(Map f, Map g, Map h)"]
- iso_char H.VV.arr_char H.VVV.arr_char H.VVV.ide_char ide_char''
+ iso_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char ide_char''
by (simp add: src_def trg_def)
qed
lemma assoc_naturality:
assumes "arr f" and "arr g" and "arr h" and "src f = trg g" and "src g = trg h"
shows "vcomp (H.HoVH (f, g, h)) (assoc (dom f) (dom g) (dom h)) =
vcomp (assoc (cod f) (cod g) (cod h)) (H.HoHV (f, g, h))"
proof -
let ?A = "Src h" and ?B = "Trg h" and ?C = "Trg g" and ?D = "Trg f"
have 1: "Src f = Trg g \<and> Src g = Trg h"
using assms src_def trg_def by simp
interpret Hom_f: category \<open>Hom ?C ?D\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_g: category \<open>Hom ?B ?C\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_h: category \<open>Hom ?A ?B\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Hom_fg: product_category \<open>Hom ?C ?D\<close> \<open>Hom ?B ?C\<close> ..
interpret Hom_gh: product_category \<open>Hom ?B ?C\<close> \<open>Hom ?A ?B\<close> ..
interpret Hom_fg_h: product_category Hom_fg.comp \<open>Hom ?A ?B\<close> ..
interpret Hom_f_gh: product_category \<open>Hom ?C ?D\<close> Hom_gh.comp ..
interpret Hom: category \<open>Hom ?A ?D\<close>
using assms ide_char arr_char dom_char cod_char category_Hom by simp
interpret Comp_fg: binary_functor \<open>Hom ?C ?D\<close> \<open>Hom ?B ?C\<close> \<open>Hom ?B ?D\<close>
\<open>\<lambda>(f', g). Comp ?D ?C ?B f' g\<close>
using assms arr_char ide_char'' binary_functor_Comp [of ?B ?C ?D] by simp
interpret Comp_gh: binary_functor \<open>Hom ?B ?C\<close> \<open>Hom ?A ?B\<close> \<open>Hom ?A ?C\<close>
\<open>\<lambda>(f', g). Comp ?C ?B ?A f' g\<close>
using assms arr_char ide_char'' binary_functor_Comp [of ?A ?B ?C] by simp
interpret Comp_fg_h: binary_functor \<open>Hom ?B ?D\<close> \<open>Hom ?A ?B\<close> \<open>Hom ?A ?D\<close>
\<open>\<lambda>(f', g). Comp ?D ?B ?A f' g\<close>
using assms arr_char ide_char'' binary_functor_Comp [of ?A ?B ?D] by simp
interpret Comp_f_gh: binary_functor \<open>Hom ?C ?D\<close> \<open>Hom ?A ?C\<close> \<open>Hom ?A ?D\<close>
\<open>\<lambda>(f', g). Comp ?D ?C ?A f' g\<close>
using assms arr_char ide_char'' binary_functor_Comp [of ?A ?C ?D] by simp
interpret \<alpha>: natural_isomorphism Hom_f_gh.comp \<open>Hom ?A ?D\<close>
\<open>\<lambda>(f', g', h'). Comp ?D ?B ?A (Comp ?D ?C ?B f' g') h'\<close>
\<open>\<lambda>(f', g', h'). Comp ?D ?C ?A f' (Comp ?C ?B ?A g' h')\<close>
\<open>\<lambda>(f', g', h'). Assoc ?D ?C ?B ?A f' g' h'\<close>
using assms ide_char arr_char natural_isomorphism_Assoc [of ?A ?B ?C ?D]
by blast
have ide_Map_dom:
"Hom_f.ide (Map (dom f)) \<and> Hom_g.ide (Map (dom g)) \<and> Hom_h.ide (Map (dom h))"
using assms 1 ide_char'' arr_char by simp
have ide_Map_cod:
"Hom_f.ide (Map (cod f)) \<and> Hom_g.ide (Map (cod g)) \<and> Hom_h.ide (Map (cod h))"
using assms 1 ide_char'' arr_char by simp
have "vcomp (H.HoVH (f, g, h)) (assoc (dom f) (dom g) (dom h)) =
vcomp (hcomp f (hcomp g h))
(MkCell (Src (dom h)) (Trg (dom f))
(Assoc (Trg (dom f)) (Trg (dom g)) (Trg (dom h)) (Src (dom h))
(Map (dom f)) (Map (dom g)) (Map (dom h))))"
- using assms 1 H.HoVH_def H.VV.arr_char H.VVV.arr_char H.VVV.ide_char
+ using assms 1 H.HoVH_def H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
assoc_def assoc_in_hom
by simp
also have "... = MkCell (Src (dom h)) (Trg (dom f))
(Hom (Src (dom h)) (Trg (dom f))
(Comp (Trg (dom f)) (Trg g) (Src (dom h)) (Map f)
(Comp (Trg g) (Trg h) (Src (dom h)) (Map g) (Map h)))
(Assoc (Trg (dom f)) (Trg (dom g)) (Trg (dom h)) (Src (dom h))
(Map (dom f)) (Map (dom g)) (Map (dom h))))"
proof -
have "Hom.seq (Comp ?D ?C ?A (Map f) (Comp ?C ?B ?A (Map g) (Map h)))
(Assoc ?D (Trg (dom g)) (Trg (dom h)) ?A
(Map (dom f)) (Map (dom g)) (Map (dom h)))"
proof (intro Hom.seqI)
show "Hom.arr (Assoc ?D (Trg (dom g)) (Trg (dom h)) ?A
(Map (dom f)) (Map (dom g)) (Map (dom h)))"
proof -
have "Hom.arr (Assoc ?D (Trg (dom g)) (Trg (dom h)) ?A
(Map (dom f)) (Map (dom g)) (Map (dom h)))"
using assms 1 arr_char Trg_dom Src_dom ide_Map_dom
\<alpha>.preserves_reflects_arr [of "(Map (dom f), Map (dom g), Map (dom h))"]
by simp
thus ?thesis
using assms 1 arr_char Src_dom Trg_dom assoc_simps(1-2) assoc_def
- H.VV.ide_char H.VV.arr_char H.VV.arr_char ide_Map_dom
+ H.VV.ide_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C ide_Map_dom
by simp
qed
show "Hom.arr (Comp ?D ?C ?A (Map f) (Comp ?C ?B ?A (Map g) (Map h)))"
proof -
have "Hom.arr (Comp ?D ?C ?A (Map f) (Comp ?C ?B ?A (Map g) (Map h)))"
using assms 1 arr_char
Comp_f_gh.preserves_reflects_arr [of "(Map f, Comp ?C ?B ?A (Map g) (Map h))"]
Comp_gh.preserves_reflects_arr [of "(Map g, Map h)"] Src_dom Trg_dom
by simp
thus ?thesis
using assms 1 arr_char Src_dom Trg_dom by simp
qed
show "Hom.dom (Comp ?D ?C ?A (Map f) (Comp ?C ?B ?A (Map g) (Map h))) =
Hom.cod (Assoc ?D (Trg (dom g)) (Trg (dom h)) ?A
(Map (dom f)) (Map (dom g)) (Map (dom h)))"
proof -
have "Hom.cod (Assoc ?D (Trg (dom g)) (Trg (dom h)) ?A
(Map (dom f)) (Map (dom g)) (Map (dom h))) =
Hom.cod (Map (assoc (dom f) (dom g) (dom h)))"
using Src_dom Trg_dom assms(1-5) assoc_simps'(5) ide_dom src_dom trg_dom
by presburger
also have "... = Comp ?D ?C ?A
(Map (dom f))
(Comp ?C ?B ?A (Map (dom g)) (Map (dom h)))"
proof -
have "dom f \<noteq> Null \<and> dom g \<noteq> Null \<and> dom h \<noteq> Null"
using arr_dom assms not_arr_null null_char by fastforce
moreover have "Hom.cod (Map (assoc (dom f) (dom g) (dom h))) =
Comp ?D ?C ?A
(Map (dom f))
(Comp ?C ?B ?A (Map (dom g)) (Map (dom h)))"
using assms assoc_simps'(5) ide_Map_dom
\<alpha>.preserves_cod [of "(Map (dom f), Map (dom g), Map (dom h))"]
by simp
ultimately show ?thesis
- using assms 1 arr_char assoc_def H.VV.ide_char H.VV.arr_char H.VV.arr_char
+ using assms 1 arr_char assoc_def H.VV.ide_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
Src_dom Trg_dom ide_Map_dom null_char assoc_simps'(5)
by simp
qed
also have "... =
Hom.dom (Comp ?D ?C ?A (Map f) (Comp ?C ?B ?A (Map g) (Map h)))"
proof -
have "arr (MkCell ?A ?C (Comp ?C ?B ?A (Map g) (Map h)))"
using assms 1 arr_char arr_MkCell
Comp_gh.preserves_reflects_arr [of "(Map g, Map h)"]
by simp
thus ?thesis
using assms arr_char 1 Map_dom
Comp_f_gh.preserves_dom [of "(Map f, Comp ?C ?B ?A (Map g) (Map h))"]
Comp_gh.preserves_dom [of "(Map g, Map h)"]
by simp
qed
finally show ?thesis by argo
qed
qed
thus ?thesis
using assms(1-5) H.hseqI' arr_char vcomp_def by auto
qed
also have "... = MkCell (Src h) (Trg f)
(Hom (Src h) (Trg f)
(Assoc (Trg f) (Trg (cod g)) (Trg (cod h)) (Src h)
(Map (cod f)) (Map (cod g)) (Map (cod h)))
(Comp (Trg f) (Trg h) (Src h)
(Comp (Trg f) (Trg g) (Src g) (Map f) (Map g)) (Map h)))"
using assms 1 arr_char \<alpha>.naturality [of "(Map f, Map g, Map h)"] by simp
also have "... = vcomp (MkCell (Src (cod h)) (Trg (cod f))
(Assoc (Trg (cod f)) (Trg (cod g)) (Trg (cod h)) (Src (cod h))
(Map (cod f)) (Map (cod g)) (Map (cod h))))
(hcomp (hcomp f g) h)"
proof -
have "arr (MkCell ?B ?D (Comp ?D ?C ?B (Map f) (Map g)))"
using assms 1 arr_char arr_MkCell
Comp_fg.preserves_reflects_arr [of "(Map f, Map g)"]
by simp
moreover have
"Hom.arr (Comp ?D ?B ?A (Comp ?D ?C ?B (Map f) (Map g)) (Map h))"
using assms 1 arr_char Comp_fg.preserves_reflects_arr [of "(Map f, Map g)"]
Comp_fg_h.preserves_reflects_arr [of "(Comp ?D ?C ?B (Map f) (Map g), Map h)"]
by simp
moreover have "Hom.arr (Assoc ?D (Trg (cod g)) (Trg (cod h)) ?A
(Map (cod f)) (Map (cod g)) (Map (cod h)))"
using assms 1 arr_char Trg_cod ide_Map_cod
\<alpha>.preserves_reflects_arr [of "(Map (cod f), Map (cod g), Map (cod h))"]
by simp
moreover have "Hom.seq (Assoc ?D (Trg (cod g)) (Trg (cod h)) ?A
(Map (cod f)) (Map (cod g)) (Map (cod h)))
(Comp ?D ?B ?A (Comp ?D ?C ?B (Map f) (Map g)) (Map h))"
proof (intro Hom.seqI)
show "Hom.arr (Comp ?D ?B ?A (Comp ?D ?C ?B (Map f) (Map g)) (Map h))"
using assms 1 arr_char Comp_fg.preserves_reflects_arr [of "(Map f, Map g)"]
Comp_fg_h.preserves_reflects_arr [of "(Comp ?D ?C ?B (Map f) (Map g), Map h)"]
by simp
show "Hom.arr (Assoc ?D (Trg (cod g)) (Trg (cod h)) ?A
(Map (cod f)) (Map (cod g)) (Map (cod h)))"
using assms 1 arr_char Trg_cod ide_Map_cod
\<alpha>.preserves_reflects_arr [of "(Map (cod f), Map (cod g), Map (cod h))"]
by simp
show "Hom.dom (Assoc ?D (Trg (cod g)) (Trg (cod h)) ?A
(Map (cod f)) (Map (cod g)) (Map (cod h))) =
Hom.cod (Comp ?D ?B ?A (Comp ?D ?C ?B (Map f) (Map g)) (Map h))"
proof -
have "Hom.dom (Assoc ?D (Trg (cod g)) (Trg (cod h)) ?A
(Map (cod f)) (Map (cod g)) (Map (cod h))) =
Hom.dom (Map (assoc (cod f) (cod g) (cod h)))"
using H.cod_trg Src_cod Trg_cod assms(1-5) assoc_simps'(5) ide_cod src_cod
trg.preserves_cod
by presburger
also have "... = Comp ?D ?B ?A
(Comp ?D ?C ?B (Map (cod f)) (Map (cod g))) (Map (cod h))"
proof -
have "cod f \<noteq> Null \<and> cod g \<noteq> Null \<and> cod h \<noteq> Null"
using arr_cod assms not_arr_null null_char by fastforce
moreover have "Hom.dom (Map (assoc (cod f) (cod g) (cod h))) =
Comp ?D ?B ?A
(Comp ?D ?C ?B (Map (cod f)) (Map (cod g))) (Map (cod h))"
using assms assoc_simps'(5) ide_Map_cod
\<alpha>.preserves_dom [of "(Map (cod f), Map (cod g), Map (cod h))"]
by simp
ultimately show ?thesis
- using assms 1 arr_char assoc_def H.VV.ide_char H.VV.arr_char H.VV.arr_char
+ using assms 1 arr_char assoc_def H.VV.ide_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
Src_cod Trg_cod ide_Map_cod null_char assoc_simps'(5)
by simp
qed
also have "... =
Hom.cod (Comp ?D ?B ?A (Comp ?D ?C ?B (Map f) (Map g)) (Map h))"
using assms arr_char 1 Map_cod arr_MkCell
Comp_fg_h.preserves_cod [of "(Comp ?D ?C ?B (Map f) (Map g), Map h)"]
Comp_fg.preserves_cod [of "(Map f, Map g)"]
Comp_fg.preserves_reflects_arr [of "(Map f, Map g)"]
by simp
finally show ?thesis by blast
qed
qed
ultimately show ?thesis
using assms arr_char vcomp_def hcomp_def Src_cod Trg_cod
- H.VV.ide_char H.VV.arr_char H.VV.arr_char src_def trg_def
+ H.VV.ide_char H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
by simp
qed
also have "... = vcomp (assoc (cod f) (cod g) (cod h)) (H.HoHV (f, g, h))"
- using assms 1 H.HoHV_def H.VV.arr_char H.VVV.arr_char H.VVV.ide_char
+ using assms 1 H.HoHV_def H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
assoc_def assoc_in_hom
by simp
finally show "vcomp (H.HoVH (f, g, h)) (assoc (dom f) (dom g) (dom h)) =
vcomp (assoc (cod f) (cod g) (cod h)) (H.HoHV (f, g, h))"
by blast
qed
interpretation \<alpha>\<^sub>0: transformation_by_components H.VVV.comp vcomp H.HoHV H.HoVH
\<open>\<lambda>(f, g, h). assoc f g h\<close>
proof
fix fgh
show "H.VVV.ide fgh \<Longrightarrow>
\<guillemotleft>case fgh of (f, g, h) \<Rightarrow> assoc f g h : H.HoHV fgh \<Rightarrow> H.HoVH fgh\<guillemotright>"
- using assoc_in_hom H.HoHV_def H.HoVH_def H.VV.arr_char H.VVV.arr_char H.VVV.ide_char
+ using assoc_in_hom H.HoHV_def H.HoVH_def H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
by (cases fgh) simp
assume fgh: "H.VVV.arr fgh"
show "vcomp (case H.VVV.cod fgh of (f, g, h) \<Rightarrow> assoc f g h) (H.HoHV fgh) =
vcomp (H.HoVH fgh) (case H.VVV.dom fgh of (f, g, h) \<Rightarrow> assoc f g h)"
using fgh assoc_simps H.HoHV_def assoc_naturality
- H.VV.arr_char H.VVV.arr_char H.VVV.dom_char H.VVV.cod_char
+ H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C H.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
by (cases fgh) simp
qed
definition \<a> ("\<a>[_,_,_]")
where "\<a> f g h == \<alpha>\<^sub>0.map (f, g, h)"
lemma \<a>_simp_ide:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "\<a>[f, g, h] =
MkCell (Src h) (Trg f)
(Assoc (Trg f) (Trg g) (Trg h) (Src h) (Map f) (Map g) (Map h))"
using assms \<a>_def assoc_def assoc_simps' MkCell_Map not_arr_null
- \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VV.arr_char H.VVV.arr_char
+ \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
interpretation \<alpha>: natural_isomorphism H.VVV.comp vcomp H.HoHV H.HoVH
\<open>\<lambda>fgh. \<a> (fst fgh) (fst (snd fgh)) (snd (snd fgh))\<close>
proof -
interpret \<alpha>: natural_isomorphism H.VVV.comp vcomp H.HoHV H.HoVH \<alpha>\<^sub>0.map
- using \<alpha>\<^sub>0.map_simp_ide iso_assoc H.VVV.ide_char H.VV.arr_char H.VVV.arr_char
+ using \<alpha>\<^sub>0.map_simp_ide iso_assoc H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
by unfold_locales auto
show "natural_isomorphism H.VVV.comp vcomp H.HoHV H.HoVH
(\<lambda>fgh. \<a> (fst fgh) (fst (snd fgh)) (snd (snd fgh)))"
using \<alpha>.natural_isomorphism_axioms \<a>_def by simp
qed
text \<open>
What remains is to show that horizontal composition with source or target defines
fully faithful functors.
\<close>
interpretation endofunctor vcomp H.L
using H.endofunctor_L by auto
interpretation endofunctor vcomp H.R
using H.endofunctor_R by auto
interpretation R: fully_faithful_functor vcomp vcomp H.R
proof
show "\<And>f f'. \<lbrakk>par f f'; H.R f = H.R f'\<rbrakk> \<Longrightarrow> f = f'"
proof -
fix \<mu> \<mu>'
assume par: "par \<mu> \<mu>'"
and eq: "H.R \<mu> = H.R \<mu>'"
have 1: "Src \<mu>' = Src \<mu> \<and> Trg \<mu>' = Trg \<mu>"
using par by (metis Src_dom Trg_dom)
interpret Hom: category \<open>Hom (Src \<mu>) (Trg \<mu>)\<close>
using par arr_char category_Hom by simp
let ?R = "\<lambda>f. if Hom.arr f
then Comp (Trg \<mu>) (Src \<mu>) (Src \<mu>) f (Id (Src \<mu>))
else Hom.null"
interpret R: fully_faithful_functor \<open>Hom (Src \<mu>) (Trg \<mu>)\<close> \<open>Hom (Src \<mu>) (Trg \<mu>)\<close> ?R
using par arr_char right_unit_Id by simp
have R\<mu>: "H.R \<mu> = MkCell (Src \<mu>) (Trg \<mu>)
(Comp (Trg \<mu>) (Src \<mu>) (Src \<mu>) (Map \<mu>) (Id (Src \<mu>)))"
unfolding hcomp_def
using par src_def null_char H.trg_src src.preserves_arr by simp
have R\<mu>': "H.R \<mu>' = MkCell (Src \<mu>) (Trg \<mu>)
(Comp (Trg \<mu>) (Src \<mu>) (Src \<mu>) (Map \<mu>') (Id (Src \<mu>)))"
unfolding hcomp_def
using par 1 src_def null_char H.trg_src src.preserves_arr by simp
have "Map \<mu> = Map \<mu>'"
using R\<mu> R\<mu>' eq par arr_char eq R.is_faithful
by (metis "1" Map_cod Map_dom cell.inject)
thus "\<mu> = \<mu>'"
using 1 par MkCell_Map by (metis arr_char null_char)
qed
show "\<And>f g \<nu>. \<lbrakk>ide f; ide g; \<guillemotleft>\<nu> : H.R f \<Rightarrow> H.R g\<guillemotright>\<rbrakk> \<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright> \<and> H.R \<mu> = \<nu>"
proof -
fix f g \<nu>
assume f: "ide f" and g: "ide g" and \<nu>: "\<guillemotleft>\<nu> : H.R f \<Rightarrow> H.R g\<guillemotright>"
interpret Hom: category \<open>Hom (Src \<nu>) (Trg \<nu>)\<close>
using \<nu> arr_char category_Hom by auto
let ?R = "\<lambda>f. if Hom.arr f
then Comp (Trg \<nu>) (Src \<nu>) (Src \<nu>) f (Id (Src \<nu>))
else Hom.null"
interpret R: fully_faithful_functor \<open>Hom (Src \<nu>) (Trg \<nu>)\<close> \<open>Hom (Src \<nu>) (Trg \<nu>)\<close> ?R
using \<nu> arr_char right_unit_Id by auto
have 0: "Src f = Src \<nu> \<and> Trg f = Trg \<nu> \<and> Src g = Src \<nu> \<and> Trg g = Trg \<nu>"
proof (intro conjI)
show "Trg f = Trg \<nu>"
using f \<nu> Trg_dom Trg_hcomp by fastforce
show "Trg g = Trg \<nu>"
using g \<nu> Trg_cod Trg_hcomp by fastforce
show "Src f = Src \<nu>"
proof -
have 1: "arr f \<and> dom f = f \<and> cod f = f"
by (meson f ide_char)
hence "hcomp f (src f) = dom \<nu>"
using \<nu> by fastforce
thus ?thesis
using 1
by (metis (no_types) H.trg_src Src.simps(1) Src_hcomp \<nu> dom_char in_homE
src.preserves_arr src_def)
qed
show "Src g = Src \<nu>"
proof -
have 1: "arr g \<and> dom g = g \<and> cod g = g"
by (meson g ide_char)
hence "hcomp g (src g) = cod \<nu>"
using \<nu> by fastforce
thus ?thesis
using 1
by (metis (no_types) H.trg_src Src.simps(1) Src_hcomp \<nu> cod_char in_homE
src.preserves_arr src_def)
qed
qed
have 1: "Hom.in_hom (Map \<nu>) (?R (Map f)) (?R (Map g))"
proof
show "Hom.arr (Map \<nu>)"
using \<nu> arr_char by auto
show "Hom.dom (Map \<nu>) = ?R (Map f)"
proof -
have 1: "arr f \<and> dom f = f \<and> cod f = f"
using f ide_char by blast
have "dom \<nu> = MkCell (Src \<nu>) (Trg \<nu>) (Hom.dom (Map \<nu>))"
by (meson \<nu> dom_char in_homE)
thus ?thesis
using 1 H.trg_src \<nu> arr_char hcomp_def src.preserves_arr src_def by fastforce
qed
show "Hom.cod (Map \<nu>) = ?R (Map g)"
proof -
have 1: "arr g \<and> dom g = g \<and> cod g = g"
using g ide_char by blast
have "cod \<nu> = MkCell (Src \<nu>) (Trg \<nu>) (Hom.cod (Map \<nu>))"
by (meson \<nu> cod_char in_homE)
thus ?thesis
using 1 H.trg_src \<nu> arr_char hcomp_def src.preserves_arr src_def by fastforce
qed
qed
have 2: "Hom.ide (Map f) \<and> Hom.ide (Map g)"
using 0 f g arr_char ide_char'' by simp
obtain x where x: "Hom.in_hom x (Map f) (Map g) \<and> ?R x = Map \<nu>"
using \<nu> 1 2 R.is_full by blast
have "\<guillemotleft>MkCell (Src \<nu>) (Trg \<nu>) x : f \<Rightarrow> g\<guillemotright>"
proof -
have "arr (MkCell (Src \<nu>) (Trg \<nu>) x)"
using \<nu> arr_char x by auto
thus ?thesis
by (metis 0 Hom.in_homE Map.simps(1) Src.simps(1) Trg.simps(1)
cod_char dom_char f g ide_char in_homI x)
qed
moreover have "H.R (MkCell (Src \<nu>) (Trg \<nu>) x) = \<nu>"
using MkCell_Map \<nu> arrI arr_char hcomp_def null_char src.preserves_arr x by auto
ultimately show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright> \<and> H.R \<mu> = \<nu>" by auto
qed
qed
interpretation L: fully_faithful_functor vcomp vcomp H.L
proof
show "\<And>f f'. \<lbrakk>par f f'; H.L f = H.L f'\<rbrakk> \<Longrightarrow> f = f'"
proof -
fix \<mu> \<mu>'
assume par: "par \<mu> \<mu>'"
and eq: "H.L \<mu> = H.L \<mu>'"
have 1: "Src \<mu>' = Src \<mu> \<and> Trg \<mu>' = Trg \<mu>"
using par by (metis Src_dom Trg_dom)
interpret Hom: category \<open>Hom (Src \<mu>) (Trg \<mu>)\<close>
using par arr_char category_Hom by simp
let ?L = "\<lambda>f. if Hom.arr f
then Comp (Trg \<mu>) (Trg \<mu>) (Src \<mu>) (Id (Trg \<mu>)) f
else Hom.null"
interpret L: fully_faithful_functor \<open>Hom (Src \<mu>) (Trg \<mu>)\<close> \<open>Hom (Src \<mu>) (Trg \<mu>)\<close> ?L
using par arr_char left_unit_Id [of "Src \<mu>" "Trg \<mu>"] by simp
have L\<mu>: "H.L \<mu> = MkCell (Src \<mu>) (Trg \<mu>)
(Comp (Trg \<mu>) (Trg \<mu>) (Src \<mu>) (Id (Trg \<mu>)) (Map \<mu>))"
unfolding hcomp_def
using par trg_def null_char H.src_trg trg.preserves_arr by simp
have L\<mu>': "H.L \<mu>' = MkCell (Src \<mu>) (Trg \<mu>)
(Comp (Trg \<mu>) (Trg \<mu>) (Src \<mu>) (Id (Trg \<mu>)) (Map \<mu>'))"
unfolding hcomp_def
using par 1 trg_def null_char H.src_trg trg.preserves_arr by simp
have "Map \<mu> = Map \<mu>'"
using L\<mu> L\<mu>' eq par arr_char eq L.is_faithful
by (metis "1" Map.simps(1) Map_cod Map_dom)
thus "\<mu> = \<mu>'"
using 1 par MkCell_Map
by (metis arr_char null_char)
qed
show "\<And>f g \<nu>. \<lbrakk>ide f; ide g; \<guillemotleft>\<nu> : H.L f \<Rightarrow> H.L g\<guillemotright>\<rbrakk> \<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright> \<and> H.L \<mu> = \<nu>"
proof -
fix f g \<nu>
assume f: "ide f" and g: "ide g" and \<nu>: "\<guillemotleft>\<nu> : H.L f \<Rightarrow> H.L g\<guillemotright>"
interpret Hom: category \<open>Hom (Src \<nu>) (Trg \<nu>)\<close>
using \<nu> arr_char category_Hom by auto
let ?L = "\<lambda>f. if Hom.arr f
then Comp (Trg \<nu>) (Trg \<nu>) (Src \<nu>) (Id (Trg \<nu>)) f
else Hom.null"
interpret L: fully_faithful_functor \<open>Hom (Src \<nu>) (Trg \<nu>)\<close> \<open>Hom (Src \<nu>) (Trg \<nu>)\<close> ?L
using \<nu> arr_char left_unit_Id by auto
have 0: "Src f = Src \<nu> \<and> Trg f = Trg \<nu> \<and> Src g = Src \<nu> \<and> Trg g = Trg \<nu>"
proof (intro conjI)
show "Src f = Src \<nu>"
using Src_dom Src_hcomp \<nu> f by fastforce
show "Src g = Src \<nu>"
using \<nu> g Src_cod Src_hcomp by fastforce
show "Trg f = Trg \<nu>"
proof -
have 1: "arr f \<and> dom f = f \<and> cod f = f"
by (meson f ide_char)
hence "hcomp (trg f) f = dom \<nu>"
using \<nu> by fastforce
thus ?thesis
using 1
by (metis (no_types) H.src_trg Trg.simps(1) Trg_hcomp \<nu> dom_char in_homE
trg.preserves_arr trg_def)
qed
show "Trg g = Trg \<nu>"
proof -
have 1: "arr g \<and> dom g = g \<and> cod g = g"
by (meson g ide_char)
hence "hcomp (trg g) g = cod \<nu>"
using \<nu> by fastforce
thus ?thesis
using 1
by (metis (no_types) H.src_trg Trg.simps(1) Trg_hcomp \<nu> cod_char in_homE
trg.preserves_arr trg_def)
qed
qed
have 1: "Hom.in_hom (Map \<nu>) (?L (Map f)) (?L (Map g))"
proof
show "Hom.arr (Map \<nu>)"
using \<nu> arr_char by auto
show "Hom.dom (Map \<nu>) = ?L (Map f)"
proof -
have "dom \<nu> = MkCell (Src \<nu>) (Trg \<nu>) (Hom.dom (Map \<nu>))"
using \<nu> dom_char [of \<nu>] by auto
hence "Hom.dom (Map \<nu>) = Map (dom \<nu>)"
by simp
also have "... = ?L (Map f)"
using 0 f \<nu> left_unit_Id [of "Src f" "Trg f"]
apply (simp add: ide_char'')
by (metis (no_types, lifting) H.src_trg Map.simps(1) Map_hcomp Trg.simps(1)
f ide_char in_homE trg.preserves_arr trg_def)
finally show ?thesis by blast
qed
show "Hom.cod (Map \<nu>) = ?L (Map g)"
proof -
have "cod \<nu> = MkCell (Src \<nu>) (Trg \<nu>) (Hom.cod (Map \<nu>))"
using \<nu> cod_char [of \<nu>] by auto
hence "Hom.cod (Map \<nu>) = Map (cod \<nu>)"
by simp
also have "... = ?L (Map g)"
using 0 g \<nu> left_unit_Id [of "Src g" "Trg g"]
apply (simp add: ide_char'')
by (metis (no_types, lifting) H.src_trg Map.simps(1) Map_hcomp Trg.simps(1)
g ide_char in_homE trg.preserves_arr trg_def)
finally show ?thesis by blast
qed
qed
have 2: "Hom.ide (Map f) \<and> Hom.ide (Map g)"
using 0 f g arr_char ide_char'' by simp
obtain x where x: "Hom.in_hom x (Map f) (Map g) \<and> ?L x = Map \<nu>"
using \<nu> 1 2 L.is_full by blast
have "\<guillemotleft>MkCell (Src \<nu>) (Trg \<nu>) x : f \<Rightarrow> g\<guillemotright>"
proof -
have "arr (MkCell (Src \<nu>) (Trg \<nu>) x)"
using \<nu> arr_char x by auto
thus ?thesis
by (metis 0 Hom.in_homE Map.simps(1) Src.simps(1) Trg.simps(1)
cod_char dom_char f g ide_char in_homI x)
qed
moreover have "H.L (MkCell (Src \<nu>) (Trg \<nu>) x) = \<nu>"
using MkCell_Map \<nu> arrI arr_char hcomp_def null_char trg.preserves_arr x by auto
ultimately show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright> \<and> H.L \<mu> = \<nu>" by auto
qed
qed
text \<open>
The unit isomorphisms are defined in terms of the specified function \<open>Unit\<close>.
\<close>
definition \<i> ("\<i>[_]")
where "\<i>[a] \<equiv> MkCell (Src a) (Src a) (Unit (Src a))"
lemma \<i>_simps [simp]:
assumes "H.obj a"
shows "Src \<i>[a] = Src a" and "Trg \<i>[a] = Trg a" and "Map \<i>[a] = Unit (Src a)"
using assms \<i>_def H.obj_def src_def trg_def
apply auto
by (metis Trg.simps(1))
text \<open>
The main result: the construction produces a bicategory.
\<close>
proposition induces_bicategory:
shows "bicategory vcomp hcomp \<a> \<i> src trg"
proof
fix a
assume a: "H.obj a"
have Src_a: "Src a \<in> Obj"
using a ide_char'' by auto
interpret Hom: category \<open>Hom (Src a) (Src a)\<close>
using Src_a category_Hom by auto
show "\<guillemotleft>\<i> a : hcomp a a \<Rightarrow> a\<guillemotright>"
proof -
have "\<guillemotleft>\<i> a : MkCell (Src a) (Src a) (Hom.dom (Unit (Src a)))
\<Rightarrow> MkCell (Src a) (Src a) (Hom.cod (Unit (Src a)))\<guillemotright>"
using a Src_a MkCell_in_hom Unit_in_hom [of "Src a"] \<i>_def
by simp (metis Hom.in_homE)
moreover have "MkCell (Src a) (Src a) (Hom.cod (Unit (Src a))) = a"
using a MkCell_Map H.obj_def Map_obj src_def Src_a Unit_in_hom by force
moreover have "MkCell (Src a) (Src a) (Hom.dom (Unit (Src a))) = hcomp a a"
using a H.obj_def Map_hcomp [of a a] Map_obj Src_a Unit_in_hom Src_hcomp Trg_hcomp
by (metis H.objE Hom.in_homE Trg.simps(1) calculation(2) hcomp_def)
ultimately show ?thesis by simp
qed
show "iso (\<i> a)"
using a Src_a iso_Unit [of "Src a"] \<open>\<guillemotleft>\<i> a : hcomp a a \<Rightarrow> a\<guillemotright>\<close> iso_char \<i>_def by auto
next
show "\<And>f g h k. \<lbrakk>ide f; ide g; ide h; ide k; src f = trg g; src g = trg h; src h = trg k\<rbrakk>
\<Longrightarrow> vcomp (hcomp f (\<a> g h k)) (vcomp (\<a> f (hcomp g h) k) (hcomp (\<a> f g h) k)) =
vcomp (\<a> f g (hcomp h k)) (\<a> (hcomp f g) h k)"
proof (intro arr_eqI)
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 1: "\<guillemotleft>hcomp f (\<a> g h k) :
hcomp f (hcomp (hcomp g h) k) \<Rightarrow> hcomp f (hcomp g (hcomp h k))\<guillemotright>"
- using f g h k fg gh hk H.VV.in_hom_char H.VV.arr_char
- assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VVV.arr_char
+ using f g h k fg gh hk H.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
H.preserves_hom \<a>_def
by auto
have 2: "\<guillemotleft>hcomp (\<a> f g h) k :
hcomp (hcomp (hcomp f g) h) k \<Rightarrow> hcomp (hcomp f (hcomp g h)) k\<guillemotright>"
- using f g h k fg gh hk H.VV.in_hom_char H.VV.arr_char
- assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VVV.arr_char
+ using f g h k fg gh hk H.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
H.preserves_hom \<a>_def
by auto
have 3: "\<guillemotleft>\<a> f (hcomp g h) k :
hcomp (hcomp f (hcomp g h)) k \<Rightarrow> hcomp f (hcomp (hcomp g h) k)\<guillemotright>"
- using f g h k fg gh hk H.VV.in_hom_char H.VV.arr_char
- assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VVV.arr_char
+ using f g h k fg gh hk H.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
H.preserves_hom \<a>_def
by auto
have 4: "seq (hcomp f (\<a> g h k)) (vcomp (\<a> f (hcomp g h) k) (hcomp (\<a> f g h) k))"
using 1 2 3 by auto
have 5: "seq (\<a> f (hcomp g h) k) (hcomp (\<a> f g h) k)"
using 2 3 by auto
have 6: "seq (\<a> f g (hcomp h k)) (\<a> (hcomp f g) h k)"
- using f g h k fg gh hk H.VV.in_hom_char H.VV.arr_char
- assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VVV.arr_char
+ using f g h k fg gh hk H.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ assoc_simps \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
H.preserves_hom \<a>_def
by simp
let ?LHS = "vcomp (hcomp f (\<a> g h k)) (vcomp (\<a> f (hcomp g h) k) (hcomp (\<a> f g h) k))"
let ?RHS = "vcomp (\<a> f g (hcomp h k)) (\<a> (hcomp f g) h k)"
show "arr ?LHS"
using 4 by simp
show "arr ?RHS"
using 6 by simp
show "Src ?LHS = Src ?RHS"
using 4 6 f g h k fg gh hk Src_vcomp Src_hcomp
- \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VV.arr_char H.VVV.arr_char
+ \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
assoc_simps assoc_simps' assoc_def \<a>_def
by simp
show "Trg ?LHS = Trg ?RHS"
using 4 6 f g h k fg gh hk Trg_vcomp Trg_hcomp
- \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char H.VV.arr_char H.VVV.arr_char
+ \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
assoc_simps assoc_simps' assoc_def \<a>_def
by simp
show "Map ?LHS = Map ?RHS"
- using 4 5 6 f g h k fg gh hk \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char
- H.VV.arr_char H.VVV.arr_char \<a>_def
+ using 4 5 6 f g h k fg gh hk \<alpha>\<^sub>0.map_simp_ide H.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
+ H.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C \<a>_def
apply simp
using Trg_src Trg_trg pentagon ideD(1) ide_char''
by (metis (no_types, lifting))
qed
qed
sublocale bicategory vcomp hcomp \<a> \<i> src trg
using induces_bicategory by simp
end
text \<open>
We now establish some correspondences between the constructed bicategory and the
originally given data, to provide some assurance that the construction really is doing
what we think it is.
\<close>
context concrete_bicategory
begin
lemma Src_in_Obj:
assumes "arr \<mu>"
shows "Src \<mu> \<in> Obj"
using assms arr_char by simp
lemma Trg_in_Obj:
assumes "arr \<mu>"
shows "Trg \<mu> \<in> Obj"
using assms arr_char by simp
lemma arr_Map:
assumes "arr \<mu>"
shows "partial_magma.arr (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>)"
using assms arr_char by simp
lemma obj_MkObj_Src:
assumes "arr \<mu>"
shows "obj (MkObj (Src \<mu>))"
using assms by (simp add: Src_in_Obj obj_MkObj)
lemma obj_MkObj_Trg:
assumes "arr \<mu>"
shows "obj (MkObj (Trg \<mu>))"
using assms by (simp add: Trg_in_Obj obj_MkObj)
lemma vcomp_MkCell [simp]:
assumes "arr (MkCell A B f)" and "arr (MkCell A B g)"
and "partial_magma.seq (Hom A B) f g"
shows "vcomp (MkCell A B f) (MkCell A B g) = MkCell A B (Hom A B f g)"
using assms(1-3) arr_char vcomp_def by fastforce
lemma hcomp_MkCell [simp]:
assumes "arr (MkCell B C f)" and "arr (MkCell A B g)"
shows "hcomp (MkCell B C f) (MkCell A B g) = MkCell A C (Comp C B A f g)"
by (simp add: assms(1-2) hcomp_def)
text \<open>
The objects of the constructed bicategory are in bijective correspondence with
the originally given set \<open>Obj\<close>, via the inverse mappings \<open>MkObj\<close> and \<open>Src\<close>.
\<close>
proposition bij_betw_obj_Obj:
shows "MkObj \<in> Obj \<rightarrow> Collect obj"
and "Src \<in> Collect obj \<rightarrow> Obj"
and "A \<in> Obj \<Longrightarrow> Src (MkObj A) = A"
and "a \<in> Collect obj \<Longrightarrow> MkObj (Src a) = a"
and "bij_betw MkObj Obj (Collect obj)"
using obj_MkObj obj_def Src_in_Obj src_def
apply auto
by (intro bij_betwI) auto
lemma obj_char:
shows "obj a \<longleftrightarrow> Src a \<in> Obj \<and> a = MkCell (Src a) (Src a) (Id (Src a))"
using Src_in_Obj bij_betw_obj_Obj(4) obj_MkObj by force
lemma Map_in_Hom:
assumes "arr \<mu>"
shows "partial_magma.in_hom (Hom (Src \<mu>) (Trg \<mu>)) (Map \<mu>) (Map (dom \<mu>)) (Map (cod \<mu>))"
by (simp add: Src_in_Obj Trg_in_Obj arr_Map assms category.in_homI category_Hom)
text \<open>
For each pair of objects \<open>a\<close> and \<open>b\<close>, the hom-category \<open>hhom a b\<close> of the constructed
bicategory is isomorphic to the originally given category \<open>Hom (Src a) (Src b)\<close>.
\<close>
proposition isomorphic_hhom_Hom:
assumes "obj a" and "obj b"
shows "isomorphic_categories
(subcategory.comp vcomp (\<lambda>f. f \<in> hhom a b)) (Hom (Src a) (Src b))"
proof -
interpret hom: subcategory vcomp \<open>\<lambda>f. f \<in> hhom a b\<close>
using assms by unfold_locales auto
interpret Hom: category \<open>Hom (Src a) (Src b)\<close>
using assms category_Hom Src_in_Obj obj_def by simp
let ?Map = "\<lambda>\<mu>. if \<mu> \<in> hhom a b then Map \<mu> else Hom.null"
let ?MkCell = "\<lambda>\<mu>. if Hom.arr \<mu> then MkCell (Src a) (Src b) \<mu> else hom.null"
interpret Map: "functor" hom.comp \<open>Hom (Src a) (Src b)\<close> ?Map
proof
fix \<mu>
show "\<not> hom.arr \<mu> \<Longrightarrow> ?Map \<mu> = Hom.null"
- using hom.inclusion [of \<mu>] hom.arr_char by auto
+ using hom.inclusion [of \<mu>] hom.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
assume \<mu>: "hom.arr \<mu>"
have 0: "src \<mu> = a \<and> trg \<mu> = b"
using \<mu> hom.arrE src_def trg_def
by (metis in_hhomE mem_Collect_eq)
have 1: "arr \<mu>"
using \<mu> hom.inclusion hom.arrE by blast
have 2: "Src \<mu> = Src a \<and> Trg \<mu> = Trg b"
using \<mu> 0
- by (metis Src_src Trg_trg hom.arr_char hom.inclusion)
+ by (metis Src_src Trg_trg hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.inclusion)
show "Hom.arr (?Map \<mu>)"
using 0 1 arr_Map [of \<mu>] \<mu> by auto
show "Hom.dom (?Map \<mu>) = ?Map (hom.dom \<mu>)"
proof -
have "Hom.dom (?Map \<mu>) = Map (dom \<mu>)"
using \<mu> hom.arrE by fastforce
thus ?thesis
by (metis \<mu> hom.arrE hom.arr_dom hom.dom_simp)
qed
show "Hom.cod (?Map \<mu>) = ?Map (hom.cod \<mu>)"
proof -
have "Hom.cod (?Map \<mu>) = Map (cod \<mu>)"
using \<mu> hom.arrE by fastforce
thus ?thesis
by (metis \<mu> hom.arrE hom.arr_cod hom.cod_simp)
qed
next
fix \<mu> \<nu>
assume \<mu>\<nu>: "hom.seq \<mu> \<nu>"
show "?Map (hom.comp \<mu> \<nu>) = Hom (Src a) (Src b) (?Map \<mu>) (?Map \<nu>)"
proof -
have 1: "hom.arr \<nu> \<and> hom.arr \<mu> \<and> seq \<mu> \<nu>"
- using \<mu>\<nu> hom.seq_char by blast
+ using \<mu>\<nu> hom.seq_char\<^sub>S\<^sub>b\<^sub>C by blast
hence 2: "hom.comp \<mu> \<nu> = vcomp \<mu> \<nu>"
using hom.comp_char by auto
have 3: "\<mu> \<in> hhom a b"
using 1 hom.arrE by blast
have "Src a = Src \<mu>"
using 3 by (metis Trg_src in_hhomE mem_Collect_eq obj_def)
moreover have "Src b = Trg \<mu>"
using 3 by (metis Trg_src Trg_trg in_hhomE mem_Collect_eq obj_def)
ultimately show ?thesis
using 1 2 Map_vcomp \<mu>\<nu> hom.arrE by presburger
qed
qed
interpret MkCell: "functor" \<open>Hom (Src a) (Src b)\<close> hom.comp ?MkCell
proof
fix \<mu>
show "\<not> Hom.arr \<mu> \<Longrightarrow> ?MkCell \<mu> = hom.null"
by simp
assume \<mu>: "Hom.arr \<mu>"
show 1: "hom.arr (?MkCell \<mu>)"
- using assms obj_def \<mu> hom.arr_char arr_MkCell src_def arr_char by auto
+ using assms obj_def \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C arr_MkCell src_def arr_char by auto
show "hom.dom (?MkCell \<mu>) = ?MkCell (Hom.dom \<mu>)"
- using assms 1 \<mu> hom.dom_char src_def arr_char obj_def Src_in_Obj by simp
+ using assms 1 \<mu> hom.dom_char\<^sub>S\<^sub>b\<^sub>C src_def arr_char obj_def Src_in_Obj by simp
show "hom.cod (?MkCell \<mu>) = ?MkCell (Hom.cod \<mu>)"
- using assms 1 \<mu> hom.cod_char src_def arr_char obj_def Src_in_Obj by simp
+ using assms 1 \<mu> hom.cod_char\<^sub>S\<^sub>b\<^sub>C src_def arr_char obj_def Src_in_Obj by simp
next
fix \<mu> \<nu>
assume \<mu>\<nu>: "Hom.seq \<mu> \<nu>"
have 1: "hom.arr (?MkCell \<mu>)"
- using assms obj_def \<mu>\<nu> hom.arr_char src_def arr_char by auto
+ using assms obj_def \<mu>\<nu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C src_def arr_char by auto
have 2: "hom.arr (?MkCell \<nu>)"
- using assms obj_def \<mu>\<nu> hom.arr_char src_def arr_char by auto
+ using assms obj_def \<mu>\<nu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C src_def arr_char by auto
have 3: "hom.dom (?MkCell \<mu>) = hom.cod (?MkCell \<nu>)"
- using \<mu>\<nu> 1 2 hom.dom_char dom_char hom.cod_char cod_char arr_char by auto
+ using \<mu>\<nu> 1 2 hom.dom_char\<^sub>S\<^sub>b\<^sub>C dom_char hom.cod_char\<^sub>S\<^sub>b\<^sub>C cod_char arr_char by auto
have 4: "seq (?MkCell \<mu>) (?MkCell \<nu>)"
by (metis 1 2 3 hom.arrE hom.cod_simp hom.comp_closed hom.dom_simp hom.inclusion)
have "hom.comp (?MkCell \<mu>) (?MkCell \<nu>) =
vcomp (MkCell (Src a) (Src b) \<mu>) (MkCell (Src a) (Src b) \<nu>)"
using \<mu>\<nu> 1 2 4 hom.comp_char by auto
also have "... = MkCell (Src a) (Src b) (Hom (Src a) (Src b) \<mu> \<nu>)"
using \<mu>\<nu> 4 vcomp_char [of "MkCell (Src a) (Src b) \<mu>" "MkCell (Src a) (Src b) \<nu>"]
by auto
also have "... = ?MkCell (Hom (Src a) (Src b) \<mu> \<nu>)"
using \<mu>\<nu> by simp
finally show "?MkCell (Hom (Src a) (Src b) \<mu> \<nu>) = hom.comp (?MkCell \<mu>) (?MkCell \<nu>)"
by simp
qed
interpret inverse_functors hom.comp \<open>Hom (Src a) (Src b)\<close> ?MkCell ?Map
proof
show "?MkCell o ?Map = hom.map"
proof
fix \<mu>
have "\<mu> \<notin> hhom a b \<Longrightarrow> (?MkCell o ?Map) \<mu> = hom.map \<mu>"
- using o_apply hom.is_extensional hom.arr_char by simp
+ using o_apply hom.is_extensional hom.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "\<mu> \<in> hhom a b \<Longrightarrow> (?MkCell o ?Map) \<mu> = hom.map \<mu>"
proof -
assume \<mu>: "\<mu> \<in> hhom a b"
have "(?MkCell o ?Map) \<mu> = MkCell (Src a) (Src b) (Map \<mu>)"
using \<mu> arr_char src_def trg_def by auto
also have "... = \<mu>"
using \<mu> MkCell_Map arr_char null_char by auto
also have "... = hom.map \<mu>"
- using \<mu> hom.arrI hom.map_def by presburger
+ using \<mu> hom.arrI\<^sub>S\<^sub>b\<^sub>C hom.map_def by presburger
finally show "(?MkCell o ?Map) \<mu> = hom.map \<mu>"
by simp
qed
ultimately show "(?MkCell o ?Map) \<mu> = hom.map \<mu>"
by blast
qed
show "?Map o ?MkCell = Hom.map"
proof
fix \<mu>
have "\<not> Hom.arr \<mu> \<Longrightarrow> (?Map o ?MkCell) \<mu> = Hom.map \<mu>"
using Hom.is_extensional hom.null_char by auto
moreover have "Hom.arr \<mu> \<Longrightarrow> (?Map o ?MkCell) \<mu> = Hom.map \<mu>"
proof -
assume \<mu>: "Hom.arr \<mu>"
have "in_hhom (MkCell (Src a) (Src b) \<mu>) a b"
- using \<mu> MkCell.preserves_reflects_arr [of \<mu>] hom.arr_char by simp
+ using \<mu> MkCell.preserves_reflects_arr [of \<mu>] hom.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
thus "(?Map o ?MkCell) \<mu> = Hom.map \<mu>"
using \<mu> by simp
qed
ultimately show "(?Map o ?MkCell) \<mu> = Hom.map \<mu>"
by blast
qed
qed
show ?thesis ..
qed
end
end
diff --git a/thys/Bicategory/EquivalenceOfBicategories.thy b/thys/Bicategory/EquivalenceOfBicategories.thy
--- a/thys/Bicategory/EquivalenceOfBicategories.thy
+++ b/thys/Bicategory/EquivalenceOfBicategories.thy
@@ -1,10421 +1,10421 @@
(* Title: EquivalenceOfBicategories
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2020
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Equivalence of Bicategories"
text \<open>
In this section, we define the notion ``equivalence of bicategories'', which generalizes
the notion of equivalence of categories, and we establish various of its properties.
In particular, we show that ``equivalent bicategories'' is an equivalence relation,
and that a pseudofunctor is part of an equivalence of bicategories if and only if it
is an equivalence pseudofunctor (\emph{i.e.}~it is faithful, locally full, locally
essentially surjective, and biessentially surjective on objects).
\<close>
theory EquivalenceOfBicategories
imports InternalAdjunction PseudonaturalTransformation
begin
subsection "Definition of Equivalence of Bicategories"
text \<open>
An equivalence of bicategories between bicategories \<open>C\<close> and \<open>D\<close> consists of pseudofunctors
\<open>F : D \<rightarrow> C\<close> and \<open>G : C \<rightarrow> D\<close> and pseudonatural equivalences \<open>\<eta>: I\<^sub>D \<approx> GF\<close> and
\<open>\<epsilon>: FG \<approx> I\<^sub>C\<close>.
\<close>
locale equivalence_of_bicategories = (* 25 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F +
G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G +
FG: composite_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G F \<Phi>\<^sub>F +
GF: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G +
I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
\<eta>: pseudonatural_equivalence V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
I\<^sub>D.map I\<^sub>D.cmp GF.map GF.cmp \<eta>\<^sub>0 \<eta>\<^sub>1 +
\<epsilon>: pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
FG.map FG.cmp I\<^sub>C.map I\<^sub>C.cmp \<epsilon>\<^sub>0 \<epsilon>\<^sub>1
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'d \<Rightarrow> 'c"
and \<Phi>\<^sub>F :: "'d * 'd \<Rightarrow> 'c"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and \<eta>\<^sub>0 :: "'d \<Rightarrow> 'd"
and \<eta>\<^sub>1 :: "'d \<Rightarrow> 'd"
and \<epsilon>\<^sub>0 :: "'c \<Rightarrow> 'c"
and \<epsilon>\<^sub>1 :: "'c \<Rightarrow> 'c"
begin
notation C.isomorphic (infix "\<cong>\<^sub>C" 50)
notation D.isomorphic (infix "\<cong>\<^sub>D" 50)
notation C.iso_in_hom ("\<guillemotleft>_ : _ \<cong>\<^sub>C _\<guillemotright>")
notation D.iso_in_hom ("\<guillemotleft>_ : _ \<cong>\<^sub>D _\<guillemotright>")
notation C.some_quasi_inverse ("_\<^sup>~\<^sup>C" [1000] 1000)
notation D.some_quasi_inverse ("_\<^sup>~\<^sup>D" [1000] 1000)
interpretation \<eta>': converse_pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
I\<^sub>D.map I\<^sub>D.cmp GF.map GF.cmp \<eta>\<^sub>0 \<eta>\<^sub>1
..
interpretation \<epsilon>': converse_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
FG.map FG.cmp I\<^sub>C.map I\<^sub>C.cmp \<epsilon>\<^sub>0 \<epsilon>\<^sub>1
..
text \<open>
Although it will be some trouble yet to prove that \<open>F\<close> is an equivalence pseudofunctor,
it is not as difficult to prove that the composites \<open>GF\<close> and \<open>FG\<close> are equivalence
pseudofunctors, by virtue of their being pseudonaturally equivalent to identity
pseudofunctors.
\<close>
interpretation GF: equivalence_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D GF.map GF.cmp
proof -
interpret GF: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
I\<^sub>D.map I\<^sub>D.cmp GF.map GF.cmp \<eta>\<^sub>0 \<eta>\<^sub>1
..
show "equivalence_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D GF.map GF.cmp"
using GF.is_equivalence_pseudofunctor by simp
qed
interpretation FG: equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C FG.map FG.cmp
proof -
interpret FG: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
I\<^sub>C.map I\<^sub>C.cmp FG.map FG.cmp \<epsilon>'.map\<^sub>0 \<epsilon>'.map\<^sub>1
..
show "equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C FG.map FG.cmp"
using FG.is_equivalence_pseudofunctor by simp
qed
text \<open>
It is also easy to establish that \<open>F\<close> and \<open>G\<close> are faithful.
We will use the facts, that \<open>GF\<close> is locally full and \<open>G\<close> is faithful,
to prove that \<open>F\<close> is locally full.
\<close>
interpretation F: faithful_functor V\<^sub>D V\<^sub>C F
proof
fix \<mu> \<mu>'
assume par: "D.par \<mu> \<mu>'" and eq: "F \<mu> = F \<mu>'"
have "src\<^sub>D \<mu> = src\<^sub>D \<mu>' \<and> trg\<^sub>D \<mu> = trg\<^sub>D \<mu>'"
using par D.src_dom D.trg_cod by (metis D.src_cod D.trg_cod)
hence "\<eta>\<^sub>0 (trg\<^sub>D \<mu>) \<star>\<^sub>D \<mu> = \<eta>\<^sub>0 (trg\<^sub>D \<mu>) \<star>\<^sub>D \<mu>'"
using par eq \<eta>.iso_map\<^sub>1_ide D.comp_arr_inv' D.trg.preserves_dom
D.comp_arr_dom D.comp_assoc
by (metis GF.is_faithful o_apply)
thus "\<mu> = \<mu>'"
using par \<eta>.ide_map\<^sub>0_obj \<eta>.components_are_equivalences
D.equivalence_cancel_left [of "\<eta>\<^sub>0 (trg\<^sub>D \<mu>)" \<mu> \<mu>']
by simp
qed
interpretation G: faithful_functor V\<^sub>C V\<^sub>D G
proof
show "\<And>\<mu> \<mu>'. \<lbrakk>C.par \<mu> \<mu>'; G \<mu> = G \<mu>'\<rbrakk> \<Longrightarrow> \<mu> = \<mu>'"
proof -
fix \<mu> \<mu>'
assume par: "C.par \<mu> \<mu>'" and eq: "G \<mu> = G \<mu>'"
have "src\<^sub>C \<mu> = src\<^sub>C \<mu>' \<and> trg\<^sub>C \<mu> = trg\<^sub>C \<mu>'"
using par by (metis C.src_cod C.trg_cod)
hence "\<mu> \<star>\<^sub>C \<epsilon>\<^sub>0 (src\<^sub>C \<mu>) = \<mu>' \<star>\<^sub>C \<epsilon>\<^sub>0 (src\<^sub>C \<mu>)"
using par eq \<epsilon>.iso_map\<^sub>1_ide C.comp_inv_arr' C.src.preserves_dom
C.comp_arr_dom C.comp_assoc
by (metis FG.is_faithful o_apply)
thus "\<mu> = \<mu>'"
using par \<epsilon>.ide_map\<^sub>0_obj \<epsilon>.components_are_equivalences
C.equivalence_cancel_right [of "\<epsilon>\<^sub>0 (src\<^sub>C \<mu>)" \<mu> \<mu>']
by simp
qed
qed
text \<open>
It is perhaps not so easy to discover a proof that \<open>F\<close> is locally essentially surjective.
Here we follow the proof of \cite{johnson-yau-2d-categories}, Lemma 6.2.13, except we
have expressed it in a way that explicitly shows its constructive nature, given that
we have already chosen an extension of each component of \<open>\<eta>\<close> and \<open>\<epsilon>\<close> to an adjoint
equivalence.
\<close>
abbreviation \<Phi>
where "\<Phi> \<psi> f f' \<equiv> \<r>\<^sub>C[f'] \<cdot>\<^sub>C
(f' \<star>\<^sub>C \<epsilon>'.counit (src\<^sub>C f)) \<cdot>\<^sub>C
\<a>\<^sub>C[f', \<epsilon>\<^sub>0 (src\<^sub>C f), \<epsilon>'.map\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>C
(C.inv (\<epsilon>\<^sub>1 f') \<star>\<^sub>C \<epsilon>'.map\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>C
\<a>\<^sub>C\<^sup>-\<^sup>1[\<epsilon>\<^sub>0 (trg\<^sub>C f), F (G f'), \<epsilon>'.map\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>C
(\<epsilon>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>C F \<psi> \<star>\<^sub>C \<epsilon>'.map\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>C
(\<epsilon>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>C C.inv (\<epsilon>'.map\<^sub>1 f)) \<cdot>\<^sub>C
\<a>\<^sub>C[\<epsilon>\<^sub>0 (trg\<^sub>C f), \<epsilon>'.map\<^sub>0 (trg\<^sub>C f), f] \<cdot>\<^sub>C
(C.inv (\<epsilon>'.counit (trg\<^sub>C f)) \<star>\<^sub>C f) \<cdot>\<^sub>C
\<l>\<^sub>C\<^sup>-\<^sup>1[f]"
lemma G_reflects_isomorphic:
assumes "C.ide f" and "C.ide f'" and "src\<^sub>C f = src\<^sub>C f'" and "trg\<^sub>C f = trg\<^sub>C f'"
and "\<guillemotleft>\<psi> : G f \<cong>\<^sub>D G f'\<guillemotright>"
shows "\<guillemotleft>\<Phi> \<psi> f f' : f \<cong>\<^sub>C f'\<guillemotright>"
proof -
let ?a = "src\<^sub>C f" and ?a' = "trg\<^sub>C f"
have f: "\<guillemotleft>f : ?a \<rightarrow>\<^sub>C ?a'\<guillemotright> \<and> C.ide f"
using assms by simp
have f': "\<guillemotleft>f' : ?a \<rightarrow>\<^sub>C ?a'\<guillemotright> \<and> C.ide f'"
using assms by simp
have \<psi>: "\<guillemotleft>\<psi> : G.map\<^sub>0 ?a \<rightarrow>\<^sub>D G.map\<^sub>0 ?a'\<guillemotright>"
using assms D.vconn_implies_hpar(1-2)
by (elim D.iso_in_homE) auto
hence F\<psi>: "\<guillemotleft>F \<psi> : F.map\<^sub>0 (G.map\<^sub>0 ?a) \<rightarrow>\<^sub>C F.map\<^sub>0 (G.map\<^sub>0 ?a')\<guillemotright>"
by auto
show "\<guillemotleft>\<Phi> \<psi> f f' : f \<cong>\<^sub>C f'\<guillemotright>"
proof (intro C.comp_iso_in_hom)
show "\<guillemotleft>\<l>\<^sub>C\<^sup>-\<^sup>1[f] : f \<cong>\<^sub>C ?a' \<star>\<^sub>C f\<guillemotright>"
using f by auto
show "\<guillemotleft>C.inv (\<epsilon>'.counit ?a') \<star>\<^sub>C f : ?a' \<star>\<^sub>C f \<cong>\<^sub>C (\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a') \<star>\<^sub>C f\<guillemotright>"
using f \<epsilon>'.counit_in_hom [of ?a'] \<epsilon>'.iso_counit [of ?a']
by (intro C.hcomp_iso_in_hom) auto
show "\<guillemotleft>\<a>\<^sub>C[\<epsilon>\<^sub>0 ?a', \<epsilon>'.map\<^sub>0 ?a', f] :
(\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a') \<star>\<^sub>C f \<cong>\<^sub>C \<epsilon>\<^sub>0 ?a' \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a' \<star>\<^sub>C f\<guillemotright>"
using f \<epsilon>'.counit_in_hom [of ?a] \<epsilon>'.ide_map\<^sub>0_obj
by (intro C.iso_in_homI) auto
show "\<guillemotleft>\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C C.inv (\<epsilon>'.map\<^sub>1 f) :
\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a' \<star>\<^sub>C f \<cong>\<^sub>C \<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F (G f) \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a\<guillemotright>"
using f \<epsilon>'.map\<^sub>1_in_hom [of f] \<epsilon>'.iso_map\<^sub>1_ide [of f] C.ide_iso_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "\<guillemotleft>\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F \<psi> \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a :
\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F (G f) \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a \<cong>\<^sub>C \<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F (G f') \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a\<guillemotright>"
using assms f f' F\<psi> F.preserves_iso C.ide_iso_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "\<guillemotleft>\<a>\<^sub>C\<^sup>-\<^sup>1[\<epsilon>\<^sub>0 ?a', F (G f'), \<epsilon>'.map\<^sub>0 ?a] :
\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F (G f') \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a \<cong>\<^sub>C (\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F (G f')) \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a\<guillemotright>"
using assms f' \<epsilon>.map\<^sub>0_in_hom(1) [of ?a'] \<epsilon>.ide_map\<^sub>0_obj [of ?a']
\<epsilon>'.map\<^sub>0_in_hom(1) [of ?a] \<epsilon>'.ide_map\<^sub>0_obj [of ?a]
C.assoc'_in_hom C.iso_assoc'
by auto
show "\<guillemotleft>C.inv (\<epsilon>\<^sub>1 f') \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a :
(\<epsilon>\<^sub>0 ?a' \<star>\<^sub>C F (G f')) \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a \<cong>\<^sub>C (f' \<star>\<^sub>C \<epsilon>\<^sub>0 ?a) \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a\<guillemotright>"
using assms f' \<epsilon>.map\<^sub>1_in_hom \<epsilon>.iso_map\<^sub>1_ide \<epsilon>'.map\<^sub>0_in_hom(1) [of ?a]
C.ide_iso_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "\<guillemotleft>\<a>\<^sub>C[f', \<epsilon>\<^sub>0 ?a, \<epsilon>'.map\<^sub>0 ?a] :
(f' \<star>\<^sub>C \<epsilon>\<^sub>0 ?a) \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a \<cong>\<^sub>C f' \<star>\<^sub>C \<epsilon>\<^sub>0 ?a \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a\<guillemotright>"
using assms f' \<epsilon>.map\<^sub>0_in_hom(1) [of ?a] \<epsilon>'.map\<^sub>0_in_hom(1) [of ?a]
\<epsilon>.ide_map\<^sub>0_obj \<epsilon>'.ide_map\<^sub>0_obj C.assoc_in_hom
by auto
show "\<guillemotleft>f' \<star>\<^sub>C \<epsilon>'.counit ?a : f' \<star>\<^sub>C \<epsilon>\<^sub>0 ?a \<star>\<^sub>C \<epsilon>'.map\<^sub>0 ?a \<cong>\<^sub>C f' \<star>\<^sub>C ?a\<guillemotright>"
using f f' \<epsilon>'.counit_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "\<guillemotleft>\<r>\<^sub>C[f'] : f' \<star>\<^sub>C ?a \<cong>\<^sub>C f'\<guillemotright>"
using assms f' by auto
qed
qed
abbreviation \<Psi>
where "\<Psi> f b b' \<equiv> \<r>\<^sub>D[G (F (\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b))] \<cdot>\<^sub>D
(G (F (\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b)) \<star>\<^sub>D \<eta>'.\<epsilon> b) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F (\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b)), \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b] \<cdot>\<^sub>D
(D.inv (\<eta>\<^sub>1 (\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b)) \<star>\<^sub>D \<eta>'.map\<^sub>0 b) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<eta>\<^sub>0 b', \<eta>'.map\<^sub>0 b', G f \<star>\<^sub>D \<eta>\<^sub>0 b] \<star>\<^sub>D \<eta>'.map\<^sub>0 b) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b', G f \<star>\<^sub>D \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b] \<cdot>\<^sub>D
((\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b]) \<cdot>\<^sub>D
(D.inv (\<eta>'.counit b') \<star>\<^sub>D G f \<star>\<^sub>D D.inv (\<eta>'.counit b)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)] \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[G f]"
lemma F_is_locally_essentially_surjective:
assumes "D.obj b" and "D.obj b'" and "\<guillemotleft>f : F.map\<^sub>0 b \<rightarrow>\<^sub>C F.map\<^sub>0 b'\<guillemotright>" and "C.ide f"
shows "\<guillemotleft>\<Phi> (\<Psi> f b b') f (F (\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b)) :
f \<cong>\<^sub>C F (\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b)\<guillemotright>"
proof -
let ?g = "\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b"
have g_in_hhom: "\<guillemotleft>?g : b \<rightarrow>\<^sub>D b'\<guillemotright>"
using assms \<eta>.ide_map\<^sub>0_obj \<eta>.map\<^sub>0_in_hhom by auto
have ide_g: "D.ide ?g"
using assms g_in_hhom \<eta>.ide_map\<^sub>0_obj \<eta>'.ide_map\<^sub>0_obj G.preserves_ide by blast
let ?\<psi> = "\<Psi> f b b'"
have "\<guillemotleft>?\<psi> : G f \<cong>\<^sub>D G (F ?g)\<guillemotright>"
proof (intro D.comp_iso_in_hom)
show "\<guillemotleft>\<r>\<^sub>D\<^sup>-\<^sup>1[G f] : G f \<cong>\<^sub>D G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)\<guillemotright>"
using assms
by (intro D.iso_in_homI) auto
show "\<guillemotleft>\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)] :
G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)
\<cong>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b') \<star>\<^sub>D G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)\<guillemotright>"
proof -
have "D.ide (G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b))"
using assms D.ide_hcomp [of "G f" "G.map\<^sub>0 (F.map\<^sub>0 b)"]
by (metis D.hcomp_in_hhomE D.hseqE D.objE F.map\<^sub>0_simps(1) G.map\<^sub>0_simps(1)
G.preserves_ide GF.map\<^sub>0_simp \<eta>.map\<^sub>0_simps(3) g_in_hhom)
thus ?thesis
using assms \<eta>.ide_map\<^sub>0_obj
by (intro D.iso_in_homI) auto
qed
show "\<guillemotleft>D.inv (\<eta>'.counit b') \<star>\<^sub>D G f \<star>\<^sub>D D.inv (\<eta>'.counit b) :
G.map\<^sub>0 (F.map\<^sub>0 b') \<star>\<^sub>D G f \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)
\<cong>\<^sub>D (\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D G f \<star>\<^sub>D (\<eta>\<^sub>0 b \<star>\<^sub>D \<eta>'.map\<^sub>0 b)\<guillemotright>"
using assms \<eta>'.iso_counit \<eta>'.counit_in_hom(2) D.vconn_implies_hpar(4) D.ide_iso_in_hom
by (intro D.hcomp_iso_in_hom) auto
show "\<guillemotleft>(\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b] :
(\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D G f \<star>\<^sub>D (\<eta>\<^sub>0 b \<star>\<^sub>D \<eta>'.map\<^sub>0 b)
\<cong>\<^sub>D (\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D (G f \<star>\<^sub>D \<eta>\<^sub>0 b) \<star>\<^sub>D \<eta>'.map\<^sub>0 b\<guillemotright>"
proof -
have "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b] :
G f \<star>\<^sub>D (\<eta>\<^sub>0 b \<star>\<^sub>D \<eta>'.map\<^sub>0 b) \<cong>\<^sub>D (G f \<star>\<^sub>D \<eta>\<^sub>0 b) \<star>\<^sub>D \<eta>'.map\<^sub>0 b\<guillemotright>"
using assms \<eta>.ide_map\<^sub>0_obj \<eta>'.ide_map\<^sub>0_obj \<eta>.map\<^sub>0_in_hom \<eta>'.map\<^sub>0_in_hom
D.assoc'_in_hom D.iso_assoc'
by (intro D.iso_in_homI) auto
thus ?thesis
using assms
by (intro D.hcomp_iso_in_hom) auto
qed
show "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b', G f \<star>\<^sub>D \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b] :
(\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D (G f \<star>\<^sub>D \<eta>\<^sub>0 b) \<star>\<^sub>D \<eta>'.map\<^sub>0 b
\<cong>\<^sub>D ((\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D (G f \<star>\<^sub>D \<eta>\<^sub>0 b)) \<star>\<^sub>D \<eta>'.map\<^sub>0 b\<guillemotright>"
using assms \<eta>.ide_map\<^sub>0_obj \<eta>'.ide_map\<^sub>0_obj \<eta>.map\<^sub>0_in_hom \<eta>'.map\<^sub>0_in_hom
D.assoc'_in_hom D.iso_assoc'
by (intro D.iso_in_homI) auto
show "\<guillemotleft>\<a>\<^sub>D[\<eta>\<^sub>0 b', \<eta>'.map\<^sub>0 b', G f \<star>\<^sub>D \<eta>\<^sub>0 b] \<star>\<^sub>D \<eta>'.map\<^sub>0 b :
((\<eta>\<^sub>0 b' \<star>\<^sub>D \<eta>'.map\<^sub>0 b') \<star>\<^sub>D (G f \<star>\<^sub>D \<eta>\<^sub>0 b)) \<star>\<^sub>D \<eta>'.map\<^sub>0 b
\<cong>\<^sub>D (\<eta>\<^sub>0 b' \<star>\<^sub>D ?g) \<star>\<^sub>D \<eta>'.map\<^sub>0 b\<guillemotright>"
using assms \<eta>.ide_map\<^sub>0_obj \<eta>'.ide_map\<^sub>0_obj \<eta>.map\<^sub>0_in_hom \<eta>'.map\<^sub>0_in_hom
D.assoc_in_hom D.iso_assoc
by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto
show "\<guillemotleft>D.inv (\<eta>\<^sub>1 ?g) \<star>\<^sub>D \<eta>'.map\<^sub>0 b :
(\<eta>\<^sub>0 b' \<star>\<^sub>D ?g) \<star>\<^sub>D \<eta>'.map\<^sub>0 b \<cong>\<^sub>D (G (F ?g) \<star>\<^sub>D \<eta>\<^sub>0 b) \<star>\<^sub>D \<eta>'.map\<^sub>0 b\<guillemotright>"
using assms \<eta>.map\<^sub>1_in_hom [of ?g] \<eta>.iso_map\<^sub>1_ide \<eta>'.map\<^sub>0_in_hom g_in_hhom ide_g
by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto
show "\<guillemotleft>\<a>\<^sub>D[G (F ?g), \<eta>\<^sub>0 b, \<eta>'.map\<^sub>0 b] :
(G (F ?g) \<star>\<^sub>D \<eta>\<^sub>0 b) \<star>\<^sub>D \<eta>'.map\<^sub>0 b \<cong>\<^sub>D G (F ?g) \<star>\<^sub>D \<eta>\<^sub>0 b \<star>\<^sub>D \<eta>'.map\<^sub>0 b\<guillemotright>"
using assms \<eta>.ide_map\<^sub>0_obj \<eta>'.ide_map\<^sub>0_obj \<eta>.map\<^sub>0_in_hom \<eta>'.map\<^sub>0_in_hom
D.assoc_in_hom D.iso_assoc
by (intro D.iso_in_homI) auto
show "\<guillemotleft>G (F ?g) \<star>\<^sub>D \<eta>'.counit b :
G (F ?g) \<star>\<^sub>D \<eta>\<^sub>0 b \<star>\<^sub>D \<eta>'.map\<^sub>0 b \<cong>\<^sub>D G (F ?g) \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)\<guillemotright>"
using assms \<eta>'.counit_in_hom D.ide_in_hom(2) ide_g
by (intro D.hcomp_iso_in_hom) auto
show "\<guillemotleft>\<r>\<^sub>D[G (F ?g)] : G (F ?g) \<star>\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b) \<cong>\<^sub>D G (F ?g)\<guillemotright>"
using assms ide_g
by (intro D.iso_in_homI) auto
qed
thus "\<guillemotleft>\<Phi> ?\<psi> f (F ?g) : f \<cong>\<^sub>C F ?g\<guillemotright>"
using assms ide_g G_reflects_isomorphic [of f "F ?g" ?\<psi>]
by (metis C.horizontal_homs_axioms D.in_hhomE F.preserves_ide F.preserves_src
F.preserves_trg g_in_hhom horizontal_homs.in_hhomE)
qed
interpretation F: equivalence_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F
proof
show "\<And>b. C.obj b \<Longrightarrow> \<exists>a. D.obj a \<and> C.equivalent_objects (F.map\<^sub>0 a) b"
by (metis FG.biessentially_surjective_on_objects FG.map\<^sub>0_simp G.map\<^sub>0_simps(1))
show "\<And>g g' \<mu>. \<lbrakk>D.ide g; D.ide g'; src\<^sub>D g = src\<^sub>D g'; trg\<^sub>D g = trg\<^sub>D g';
\<guillemotleft>\<mu> : F g \<Rightarrow>\<^sub>C F g'\<guillemotright>\<rbrakk>
\<Longrightarrow> \<exists>\<nu>. \<guillemotleft>\<nu> : g \<Rightarrow>\<^sub>D g'\<guillemotright> \<and> F \<nu> = \<mu>"
proof -
fix g g' \<mu>
assume g: "D.ide g" and g': "D.ide g'"
assume eq_src: "src\<^sub>D g = src\<^sub>D g'" and eq_trg: "trg\<^sub>D g = trg\<^sub>D g'"
assume \<mu>: "\<guillemotleft>\<mu> : F g \<Rightarrow>\<^sub>C F g'\<guillemotright>"
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu> : g \<Rightarrow>\<^sub>D g'\<guillemotright> \<and> G (F \<nu>) = G \<mu>"
using g g' eq_src eq_trg \<mu> GF.locally_full [of g g' "G \<mu>"] by auto
have "F \<nu> = \<mu>"
using \<mu> \<nu> G.is_faithful
by (metis (mono_tags, lifting) C.in_homE F.preserves_hom)
thus "\<exists>\<nu>. \<guillemotleft>\<nu> : g \<Rightarrow>\<^sub>D g'\<guillemotright> \<and> F \<nu> = \<mu>"
using \<nu> by auto
qed
show "\<And>b b' f. \<lbrakk>D.obj b; D.obj b'; \<guillemotleft>f : F.map\<^sub>0 b \<rightarrow>\<^sub>C F.map\<^sub>0 b'\<guillemotright>; C.ide f\<rbrakk>
\<Longrightarrow> \<exists>g. \<guillemotleft>g : b \<rightarrow>\<^sub>D b'\<guillemotright> \<and> D.ide g \<and> C.isomorphic (F g) f"
proof -
fix b b' f
assume b: "D.obj b" and b': "D.obj b'" and f: "\<guillemotleft>f : F.map\<^sub>0 b \<rightarrow>\<^sub>C F.map\<^sub>0 b'\<guillemotright>"
assume ide_f: "C.ide f"
let ?g = "\<eta>'.map\<^sub>0 b' \<star>\<^sub>D G f \<star>\<^sub>D \<eta>\<^sub>0 b"
have g_in_hhom: "\<guillemotleft>?g : b \<rightarrow>\<^sub>D b'\<guillemotright>"
using b b' f \<eta>.ide_map\<^sub>0_obj \<eta>.map\<^sub>0_in_hhom by auto
have ide_g: "D.ide ?g"
using b b' f ide_f g_in_hhom \<eta>.ide_map\<^sub>0_obj \<eta>'.ide_map\<^sub>0_obj G.preserves_ide by blast
have "f \<cong>\<^sub>C F ?g"
using b b' f ide_f F_is_locally_essentially_surjective C.isomorphic_symmetric
by (meson C.isomorphicI')
thus "\<exists>g. \<guillemotleft>g : b \<rightarrow>\<^sub>D b'\<guillemotright> \<and> D.ide g \<and> F g \<cong>\<^sub>C f"
using g_in_hhom ide_g C.isomorphic_symmetric C.isomorphic_def by blast
qed
qed
lemma equivalence_pseudofunctor_left:
shows "equivalence_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F"
..
end
subsection "Equivalences Respect Pseudonatural Equivalence"
text \<open>
In this section we prove that, in an equivalence of bicategories, either pseudofunctor
may be replaced by a pseudonaturally equivalent one and still obtain an equivalence of
bicategories.
\<close>
locale equivalence_of_bicategories_and_pseudonatural_equivalence_left = (* 30 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
E: equivalence_of_bicategories +
\<tau>: pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F F' \<Phi>\<^sub>F' \<tau>\<^sub>0 \<tau>\<^sub>1
for F'
and \<Phi>\<^sub>F'
and \<tau>\<^sub>0
and \<tau>\<^sub>1
(*
* TODO: If I attempt to declare these free variables with the types they are ultimately
* inferred to have, then a disjoint set of type variables gets used, resulting in an error.
*)
begin
interpretation GF': composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F' \<Phi>\<^sub>F' G \<Phi>\<^sub>G
..
interpretation F'G: composite_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G F' \<Phi>\<^sub>F'
..
interpretation \<tau>': converse_pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F F' \<Phi>\<^sub>F' \<tau>\<^sub>0 \<tau>\<^sub>1
..
interpretation \<tau>'oG: pseudonatural_equivalence_whisker_right V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F' \<Phi>\<^sub>F' F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>'.map\<^sub>0 \<tau>'.map\<^sub>1
..
interpretation Go\<tau>: pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi>\<^sub>F F' \<Phi>\<^sub>F' G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
..
sublocale unit: composite_pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
E.I\<^sub>D.map E.I\<^sub>D.cmp E.GF.map E.GF.cmp GF'.map GF'.cmp
\<eta>\<^sub>0 \<eta>\<^sub>1 Go\<tau>.map\<^sub>0 Go\<tau>.map\<^sub>1
..
sublocale counit: composite_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F'G.map F'G.cmp E.FG.map E.FG.cmp E.I\<^sub>C.map E.I\<^sub>C.cmp
\<tau>'oG.map\<^sub>0 \<tau>'oG.map\<^sub>1 \<epsilon>\<^sub>0 \<epsilon>\<^sub>1
..
definition unit\<^sub>0
where "unit\<^sub>0 \<equiv> unit.map\<^sub>0"
definition unit\<^sub>1
where "unit\<^sub>1 \<equiv> unit.map\<^sub>1"
definition counit\<^sub>0
where "counit\<^sub>0 \<equiv> counit.map\<^sub>0"
definition counit\<^sub>1
where "counit\<^sub>1 \<equiv> counit.map\<^sub>1"
sublocale equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F' \<Phi>\<^sub>F' G \<Phi>\<^sub>G
unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1
unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def
..
lemma induces_equivalence_of_bicategories:
shows "equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F' \<Phi>\<^sub>F' G \<Phi>\<^sub>G
unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
..
end
locale equivalence_of_bicategories_and_pseudonatural_equivalence_right = (* 30 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
E: equivalence_of_bicategories +
\<tau>: pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G G' \<Phi>\<^sub>G' \<tau>\<^sub>0 \<tau>\<^sub>1
for G'
and \<Phi>\<^sub>G'
and \<tau>\<^sub>0
and \<tau>\<^sub>1
begin
interpretation G'F: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G' \<Phi>\<^sub>G'
..
interpretation FG': composite_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G' \<Phi>\<^sub>G' F \<Phi>\<^sub>F
..
interpretation \<tau>': converse_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G G' \<Phi>\<^sub>G' \<tau>\<^sub>0 \<tau>\<^sub>1
..
interpretation \<tau>oF: pseudonatural_equivalence_whisker_right V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
G \<Phi>\<^sub>G G' \<Phi>\<^sub>G' F \<Phi>\<^sub>F \<tau>\<^sub>0 \<tau>\<^sub>1
..
interpretation Fo\<tau>': pseudonatural_equivalence_whisker_left
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
G' \<Phi>\<^sub>G' G \<Phi>\<^sub>G F \<Phi>\<^sub>F \<tau>'.map\<^sub>0 \<tau>'.map\<^sub>1
..
sublocale unit: composite_pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
E.I\<^sub>D.map E.I\<^sub>D.cmp E.GF.map E.GF.cmp G'F.map G'F.cmp
\<eta>\<^sub>0 \<eta>\<^sub>1 \<tau>oF.map\<^sub>0 \<tau>oF.map\<^sub>1
..
sublocale counit: composite_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
FG'.map FG'.cmp E.FG.map E.FG.cmp E.I\<^sub>C.map E.I\<^sub>C.cmp
Fo\<tau>'.map\<^sub>0 Fo\<tau>'.map\<^sub>1 \<epsilon>\<^sub>0 \<epsilon>\<^sub>1
..
definition unit\<^sub>0
where "unit\<^sub>0 \<equiv> unit.map\<^sub>0"
definition unit\<^sub>1
where "unit\<^sub>1 \<equiv> unit.map\<^sub>1"
definition counit\<^sub>0
where "counit\<^sub>0 \<equiv> counit.map\<^sub>0"
definition counit\<^sub>1
where "counit\<^sub>1 \<equiv> counit.map\<^sub>1"
sublocale equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G' \<Phi>\<^sub>G'
unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1
unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def
..
lemma induces_equivalence_of_bicategories:
shows "equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G' \<Phi>\<^sub>G'
unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
..
end
subsection "Converse of an Equivalence"
text \<open>
Equivalence of bicategories is a symmetric notion, in the sense that from an equivalence
of bicategories from \<open>C\<close> to \<open>D\<close> we may obtain an equivalence of bicategories from
\<open>D\<close> to \<open>C\<close>. The converse equivalence is obtained by interchanging the pseudofunctors
\<open>F\<close> and \<open>G\<close> and replacing the pseudonatural equivalences \<open>\<eta>\<close> and \<open>\<epsilon>\<close> by converse
equivalences. Essentially all the work goes into proving that pseudonatural equivalences
have pseudonatural converses, which we have already done.
\<close>
locale converse_equivalence_of_bicategories = (* 25 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
E: equivalence_of_bicategories
begin
sublocale counit: converse_pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
I\<^sub>D.map I\<^sub>D.cmp E.GF.map E.GF.cmp \<eta>\<^sub>0 \<eta>\<^sub>1
..
sublocale unit: converse_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
E.FG.map E.FG.cmp I\<^sub>C.map I\<^sub>C.cmp \<epsilon>\<^sub>0 \<epsilon>\<^sub>1
..
definition unit\<^sub>0
where "unit\<^sub>0 \<equiv> unit.map\<^sub>0"
definition unit\<^sub>1
where "unit\<^sub>1 \<equiv> unit.map\<^sub>1"
definition counit\<^sub>0
where "counit\<^sub>0 \<equiv> counit.map\<^sub>0"
definition counit\<^sub>1
where "counit\<^sub>1 \<equiv> counit.map\<^sub>1"
sublocale equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G F \<Phi>\<^sub>F
unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1
unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def
..
lemma is_equivalence_of_bicategories:
shows "equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G F \<Phi>\<^sub>F
unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
..
end
subsection "Composition of Equivalences"
text \<open>
An equivalence of bicategories from \<open>B\<close> to \<open>C\<close> and an equivalence of bicategories
from \<open>C\<close> to \<open>D\<close> may be composed to obtain an equivalence of bicategories
from \<open>B\<close> to \<open>D\<close>.
\<close>
locale composite_equivalence_of_bicategories = (* 60 sec *)
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
I\<^sub>B: identity_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F_: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B F \<Phi>\<^sub>F +
G_: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G +
H: pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C H \<Phi>\<^sub>H +
K: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D K \<Phi>\<^sub>K +
FG: equivalence_of_bicategories
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<rho>\<^sub>0 \<rho>\<^sub>1 \<sigma>\<^sub>0 \<sigma>\<^sub>1 +
HK: equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H K \<Phi>\<^sub>K \<zeta>\<^sub>0 \<zeta>\<^sub>1 \<xi>\<^sub>0 \<xi>\<^sub>1
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
and V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'b"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'b"
and G :: "'b \<Rightarrow> 'c"
and \<Phi>\<^sub>G :: "'b * 'b \<Rightarrow> 'c"
and H :: "'d \<Rightarrow> 'c"
and \<Phi>\<^sub>H :: "'d * 'd \<Rightarrow> 'c"
and K :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>K :: "'c * 'c \<Rightarrow> 'd"
and \<rho>\<^sub>0 :: "'c \<Rightarrow> 'c"
and \<rho>\<^sub>1 :: "'c \<Rightarrow> 'c"
and \<sigma>\<^sub>0 :: "'b \<Rightarrow> 'b"
and \<sigma>\<^sub>1 :: "'b \<Rightarrow> 'b"
and \<zeta>\<^sub>0 :: "'d \<Rightarrow> 'd"
and \<zeta>\<^sub>1 :: "'d \<Rightarrow> 'd"
and \<xi>\<^sub>0 :: "'c \<Rightarrow> 'c"
and \<xi>\<^sub>1 :: "'c \<Rightarrow> 'c"
begin
notation B.\<a>' ("\<a>\<^sub>B\<^sup>-\<^sup>1[_, _, _]")
text \<open>
At this point we could make the explicit definitions:
\begin{itemize}
\item \<open>\<eta>\<^sub>0 = K (\<rho>\<^sub>0 (H.map\<^sub>0 a)) \<star>\<^sub>D \<zeta>\<^sub>0 a\<close>
\item \<open>\<eta>\<^sub>1 = \<a>\<^sub>D\<^sup>-\<^sup>1[K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), \<zeta>\<^sub>0 (trg\<^sub>D f), f] \<cdot>\<^sub>D
(K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))) \<star>\<^sub>D \<zeta>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), K (H f), \<zeta>\<^sub>0 (src\<^sub>D f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f)), H f)) \<cdot>\<^sub>D
K (\<rho>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<Phi>\<^sub>K (G (F (H f)), \<rho>\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))) \<star>\<^sub>D \<zeta>\<^sub>0 (src\<^sub>D f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[K (G (F (H f))), K (\<rho>\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))), \<zeta>\<^sub>0 (src\<^sub>D f)]\<close>
\item \<open>\<epsilon>\<^sub>0 = \<sigma>\<^sub>0 a \<star>\<^sub>B F (\<xi>\<^sub>0 (G_.map\<^sub>0 a))\<close>
\item \<open>\<epsilon>\<^sub>1 = \<a>\<^sub>B\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>B f), F (\<xi>\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f))), F (H (K (G f)))] \<cdot>\<^sub>B
(\<sigma>\<^sub>0 (trg\<^sub>B f) \<star>\<^sub>B
B.inv (\<Phi>\<^sub>F (\<xi>\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f)), H (K (G f)))) \<cdot>\<^sub>B
F (\<xi>\<^sub>1 (G f)) \<cdot>\<^sub>B \<Phi>\<^sub>F (G f, \<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>B
\<a>\<^sub>B[\<sigma>\<^sub>0 (trg\<^sub>B f), F (G f), F (\<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))] \<cdot>\<^sub>B
(\<sigma>\<^sub>1 f \<star>\<^sub>B F (\<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, \<sigma>\<^sub>0 (src\<^sub>B f), F (\<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))]\<close>
\end{itemize}
but then it is a daunting task to establish the necessary coherence conditions.
It is easier (and more useful) to use general results about composite pseudonatural
equivalences, which are somewhat easier to prove, though long calculations are
still required for those.
\<close>
sublocale FH: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B H \<Phi>\<^sub>H F \<Phi>\<^sub>F
..
sublocale KG: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G K \<Phi>\<^sub>K
..
interpretation IG: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
G \<Phi>\<^sub>G I\<^sub>C.map I\<^sub>C.cmp
..
interpretation IH: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
H \<Phi>\<^sub>H I\<^sub>C.map I\<^sub>C.cmp
..
interpretation HKG: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
G \<Phi>\<^sub>G HK.FG.map HK.FG.cmp
..
interpretation GFH: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
H \<Phi>\<^sub>H FG.GF.map FG.GF.cmp
..
interpretation KGFH: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
GFH.map GFH.cmp K \<Phi>\<^sub>K
..
interpretation FHKG: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
HKG.map HKG.cmp F \<Phi>\<^sub>F
..
interpretation \<rho>oH: pseudonatural_equivalence_whisker_right V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
I\<^sub>C.map I\<^sub>C.cmp FG.GF.map FG.GF.cmp H \<Phi>\<^sub>H \<rho>\<^sub>0 \<rho>\<^sub>1
..
interpretation Ko\<rho>oH: pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
H \<Phi>\<^sub>H GFH.map GFH.cmp K \<Phi>\<^sub>K \<rho>oH.map\<^sub>0 \<rho>oH.map\<^sub>1
proof -
interpret Ko\<rho>oH: pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
IH.map IH.cmp GFH.map GFH.cmp K \<Phi>\<^sub>K \<rho>oH.map\<^sub>0 \<rho>oH.map\<^sub>1
..
have "IH.map = H"
using H.is_extensional IH.is_extensional H.functor_axioms by force
moreover have "IH.cmp = \<Phi>\<^sub>H"
proof
fix \<mu>\<nu>
show "IH.cmp \<mu>\<nu> = \<Phi>\<^sub>H \<mu>\<nu>"
- using IH.cmp_def D.VV.arr_char D.VV.dom_simp D.VV.cod_simp H.FF_def
+ using IH.cmp_def D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp H.FF_def
C.comp_arr_dom C.comp_cod_arr H.\<Phi>.is_natural_1 H.\<Phi>.is_extensional
by (cases "D.VV.arr \<mu>\<nu>") auto
qed
ultimately show "pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
H \<Phi>\<^sub>H GFH.map GFH.cmp K \<Phi>\<^sub>K \<rho>oH.map\<^sub>0 \<rho>oH.map\<^sub>1"
using Ko\<rho>oH.pseudonatural_equivalence_whisker_left_axioms by simp
qed
interpretation \<xi>oG: pseudonatural_equivalence_whisker_right V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
HK.FG.map HK.FG.cmp I\<^sub>C.map I\<^sub>C.cmp G \<Phi>\<^sub>G \<xi>\<^sub>0 \<xi>\<^sub>1
..
interpretation Fo\<xi>oG: pseudonatural_equivalence_whisker_left V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
HKG.map HKG.cmp G \<Phi>\<^sub>G F \<Phi>\<^sub>F \<xi>oG.map\<^sub>0 \<xi>oG.map\<^sub>1
proof -
interpret Fo\<xi>oG: pseudonatural_equivalence_whisker_left V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
HKG.map HKG.cmp IG.map IG.cmp F \<Phi>\<^sub>F \<xi>oG.map\<^sub>0 \<xi>oG.map\<^sub>1
..
have "IG.map = G"
using G_.is_extensional IG.is_extensional
by (meson G_.functor_axioms comp_identity_functor)
moreover have "IG.cmp = \<Phi>\<^sub>G"
proof
fix \<mu>\<nu>
show "IG.cmp \<mu>\<nu> = \<Phi>\<^sub>G \<mu>\<nu>"
- using IG.cmp_def B.VV.arr_char B.VV.dom_simp B.VV.cod_simp G_.FF_def
+ using IG.cmp_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp G_.FF_def
C.comp_arr_dom C.comp_cod_arr G_.\<Phi>.is_natural_1 G_.\<Phi>.is_extensional
by (cases "B.VV.arr \<mu>\<nu>") auto
qed
ultimately show "pseudonatural_equivalence_whisker_left V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
HKG.map HKG.cmp G \<Phi>\<^sub>G F \<Phi>\<^sub>F \<xi>oG.map\<^sub>0 \<xi>oG.map\<^sub>1"
using Fo\<xi>oG.pseudonatural_equivalence_whisker_left_axioms by simp
qed
sublocale unit: composite_pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
I\<^sub>D.map I\<^sub>D.cmp HK.GF.map HK.GF.cmp KGFH.map KGFH.cmp
\<zeta>\<^sub>0 \<zeta>\<^sub>1 Ko\<rho>oH.map\<^sub>0 Ko\<rho>oH.map\<^sub>1
..
sublocale counit: composite_pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
FHKG.map FHKG.cmp FG.FG.map FG.FG.cmp I\<^sub>B.map I\<^sub>B.cmp
Fo\<xi>oG.map\<^sub>0 Fo\<xi>oG.map\<^sub>1 \<sigma>\<^sub>0 \<sigma>\<^sub>1
..
abbreviation left_map
where "left_map \<equiv> FH.map"
abbreviation right_map
where "right_map \<equiv> KG.map"
abbreviation left_cmp
where "left_cmp \<equiv> FH.cmp"
abbreviation right_cmp
where "right_cmp \<equiv> KG.cmp"
definition unit\<^sub>0
where "unit\<^sub>0 \<equiv> unit.map\<^sub>0"
definition unit\<^sub>1
where "unit\<^sub>1 \<equiv> unit.map\<^sub>1"
definition counit\<^sub>0
where "counit\<^sub>0 \<equiv> counit.map\<^sub>0"
definition counit\<^sub>1
where "counit\<^sub>1 == counit.map\<^sub>1"
lemma unit\<^sub>0_simp:
assumes "D.obj a"
shows "unit\<^sub>0 a = K (\<rho>\<^sub>0 (H.map\<^sub>0 a)) \<star>\<^sub>D \<zeta>\<^sub>0 a"
using assms unit\<^sub>0_def unit.map\<^sub>0_def Ko\<rho>oH.map\<^sub>0_def \<rho>oH.map\<^sub>0_def by simp
lemma unit\<^sub>1_simp:
assumes "D.ide f"
shows "unit\<^sub>1 f = \<a>\<^sub>D\<^sup>-\<^sup>1[K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), \<zeta>\<^sub>0 (trg\<^sub>D f), f] \<cdot>\<^sub>D
(K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))) \<star>\<^sub>D \<zeta>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), K (H f), \<zeta>\<^sub>0 (src\<^sub>D f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>K (\<rho>\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f)), H f)) \<cdot>\<^sub>D
K (\<rho>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<Phi>\<^sub>K (G (F (H f)), \<rho>\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))) \<star>\<^sub>D \<zeta>\<^sub>0 (src\<^sub>D f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[K (G (F (H f))), K (\<rho>\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))), \<zeta>\<^sub>0 (src\<^sub>D f)]"
using assms unit\<^sub>1_def unit.map\<^sub>1_def Ko\<rho>oH.map\<^sub>0_def \<rho>oH.map\<^sub>0_def Ko\<rho>oH.map\<^sub>1_def \<rho>oH.map\<^sub>1_def
by simp
lemma counit\<^sub>0_simp:
assumes "B.obj a"
shows "counit\<^sub>0 a = \<sigma>\<^sub>0 a \<star>\<^sub>B F (\<xi>\<^sub>0 (G_.map\<^sub>0 a))"
using assms counit\<^sub>0_def counit.map\<^sub>0_def Fo\<xi>oG.map\<^sub>0_def \<xi>oG.map\<^sub>0_def by simp
lemma counit\<^sub>1_simp:
assumes "B.ide f"
shows "counit\<^sub>1 f = \<a>\<^sub>B\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>B f), F (\<xi>\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f))), F (H (K (G f)))] \<cdot>\<^sub>B
(\<sigma>\<^sub>0 (trg\<^sub>B f) \<star>\<^sub>B
B.inv (\<Phi>\<^sub>F (\<xi>\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f)), H (K (G f)))) \<cdot>\<^sub>B
F (\<xi>\<^sub>1 (G f)) \<cdot>\<^sub>B
\<Phi>\<^sub>F (G f, \<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>B
\<a>\<^sub>B[\<sigma>\<^sub>0 (trg\<^sub>B f), F (G f), F (\<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))] \<cdot>\<^sub>B
(\<sigma>\<^sub>1 f \<star>\<^sub>B F (\<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, \<sigma>\<^sub>0 (src\<^sub>B f), F (\<xi>\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))]"
using assms counit\<^sub>1_def counit.map\<^sub>1_def Fo\<xi>oG.map\<^sub>0_def Fo\<xi>oG.map\<^sub>1_def
\<xi>oG.map\<^sub>0_def \<xi>oG.map\<^sub>1_def
by simp
sublocale equivalence_of_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
FH.map FH.cmp KG.map KG.cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1
proof -
interpret FH_KG: composite_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
KG.map KG.cmp FH.map FH.cmp
..
interpret KG_FH: composite_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
FH.map FH.cmp KG.map KG.cmp
..
have "FH_KG.map = FHKG.map" by auto
moreover have "FH_KG.cmp = FHKG.cmp"
proof
fix \<mu>\<nu>
show "FH_KG.cmp \<mu>\<nu> = FHKG.cmp \<mu>\<nu>"
proof (cases "B.VV.arr \<mu>\<nu>")
case False
thus ?thesis
using FH_KG.\<Phi>.is_extensional FHKG.\<Phi>.is_extensional by simp
next
case True
have "FH_KG.cmp \<mu>\<nu> =
F (H (K (G (I\<^sub>B.cmp \<mu>\<nu>)))) \<cdot>\<^sub>B
F (H (K (G (B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (K (\<Phi>\<^sub>G (B.dom (fst \<mu>\<nu>), B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (\<Phi>\<^sub>K (G (B.dom (fst \<mu>\<nu>)), G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
(F (H (K (G (B.dom (fst \<mu>\<nu>))) \<star>\<^sub>D K (G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (\<Phi>\<^sub>H (K (G (B.dom (fst \<mu>\<nu>))), K (G (B.dom (snd \<mu>\<nu>)))))) \<cdot>\<^sub>B
\<Phi>\<^sub>F (H (K (G (B.dom (fst \<mu>\<nu>)))), H (K (G (B.dom (snd \<mu>\<nu>)))))"
using True FH_KG.cmp_def FHKG.cmp_def HKG.cmp_def
FH.cmp_def HK.FG.cmp_def KG.cmp_def
- B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
G_.FF_def H.FF_def K.FF_def B.comp_assoc
by auto
also have "... = F (H (K (G (I\<^sub>B.cmp \<mu>\<nu>)))) \<cdot>\<^sub>B
F (H (K (G (B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (K (\<Phi>\<^sub>G (B.dom (fst \<mu>\<nu>), B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (\<Phi>\<^sub>K (G (B.dom (fst \<mu>\<nu>)), G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (\<Phi>\<^sub>H (K (G (B.dom (fst \<mu>\<nu>))), K (G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
\<Phi>\<^sub>F (H (K (G (B.dom (fst \<mu>\<nu>)))), H (K (G (B.dom (snd \<mu>\<nu>)))))"
proof -
have "F (H (K (G (B.dom (fst \<mu>\<nu>))) \<star>\<^sub>D K (G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (\<Phi>\<^sub>H (K (G (B.dom (fst \<mu>\<nu>))), K (G (B.dom (snd \<mu>\<nu>))))) =
F (\<Phi>\<^sub>H (K (G (B.dom (fst \<mu>\<nu>))), K (G (B.dom (snd \<mu>\<nu>)))))"
- using True B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ using True B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
B.comp_cod_arr
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = F (H (K (G (I\<^sub>B.cmp \<mu>\<nu>)))) \<cdot>\<^sub>B
F (H (K (G (B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (K (\<Phi>\<^sub>G (B.dom (fst \<mu>\<nu>), B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
(F (H (K (G (B.dom (fst \<mu>\<nu>)) \<star>\<^sub>C G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (\<Phi>\<^sub>K (G (B.dom (fst \<mu>\<nu>)), G (B.dom (snd \<mu>\<nu>)))))) \<cdot>\<^sub>B
F (\<Phi>\<^sub>H (K (G (B.dom (fst \<mu>\<nu>))), K (G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
\<Phi>\<^sub>F (H (K (G (B.dom (fst \<mu>\<nu>)))), H (K (G (B.dom (snd \<mu>\<nu>)))))"
proof -
have "F (H (K (G (B.dom (fst \<mu>\<nu>)) \<star>\<^sub>C G (B.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>B
F (H (\<Phi>\<^sub>K (G (B.dom (fst \<mu>\<nu>)), G (B.dom (snd \<mu>\<nu>))))) =
F (H (\<Phi>\<^sub>K (G (B.dom (fst \<mu>\<nu>)), G (B.dom (snd \<mu>\<nu>)))))"
- using True B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ using True B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
G_.FF_def H.FF_def K.FF_def B.comp_assoc
B.comp_cod_arr
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = FHKG.cmp \<mu>\<nu>"
using True FH_KG.cmp_def FHKG.cmp_def HKG.cmp_def
FH.cmp_def HK.FG.cmp_def KG.cmp_def
- B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
G_.FF_def H.FF_def K.FF_def B.comp_assoc
by simp
finally show ?thesis by blast
qed
qed
moreover have "KG_FH.map = KGFH.map" by auto
moreover have "KG_FH.cmp = KGFH.cmp"
proof
fix \<mu>\<nu>
show "KG_FH.cmp \<mu>\<nu> = KGFH.cmp \<mu>\<nu>"
proof (cases "D.VV.arr \<mu>\<nu>")
case False
thus ?thesis
using KG_FH.\<Phi>.is_extensional KGFH.\<Phi>.is_extensional by simp
next
case True
have "KG_FH.cmp \<mu>\<nu> =
K (G (F (H (I\<^sub>D.cmp \<mu>\<nu>)))) \<cdot>\<^sub>D
K (G (F (H (D.dom (fst \<mu>\<nu>) \<star>\<^sub>D D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (F (\<Phi>\<^sub>H (D.dom (fst \<mu>\<nu>), D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (\<Phi>\<^sub>F (H (D.dom (fst \<mu>\<nu>)), H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
(K (G (F (H (D.dom (fst \<mu>\<nu>))) \<star>\<^sub>B F (H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (\<Phi>\<^sub>G (F (H (D.dom (fst \<mu>\<nu>))), F (H (D.dom (snd \<mu>\<nu>)))))) \<cdot>\<^sub>D
\<Phi>\<^sub>K (G (F (H (D.dom (fst \<mu>\<nu>)))), G (F (H (D.dom (snd \<mu>\<nu>)))))"
using True KG_FH.cmp_def FH.cmp_def KG.cmp_def
- B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
G_.FF_def H.FF_def K.FF_def
D.comp_assoc
by auto
also have "... = K (G (F (H (I\<^sub>D.cmp \<mu>\<nu>)))) \<cdot>\<^sub>D
K (G (F (H (D.dom (fst \<mu>\<nu>) \<star>\<^sub>D D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (F (\<Phi>\<^sub>H (D.dom (fst \<mu>\<nu>), D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (\<Phi>\<^sub>F (H (D.dom (fst \<mu>\<nu>)), H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (\<Phi>\<^sub>G (F (H (D.dom (fst \<mu>\<nu>))), F (H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
\<Phi>\<^sub>K (G (F (H (D.dom (fst \<mu>\<nu>)))), G (F (H (D.dom (snd \<mu>\<nu>)))))"
proof -
have "K (G (F (H (D.dom (fst \<mu>\<nu>))) \<star>\<^sub>B F (H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (\<Phi>\<^sub>G (F (H (D.dom (fst \<mu>\<nu>))), F (H (D.dom (snd \<mu>\<nu>))))) =
K (\<Phi>\<^sub>G (F (H (D.dom (fst \<mu>\<nu>))), F (H (D.dom (snd \<mu>\<nu>)))))"
using True KG_FH.cmp_def FH.cmp_def KG.cmp_def
- B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
D.comp_cod_arr
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = K (G (F (H (I\<^sub>D.cmp \<mu>\<nu>)))) \<cdot>\<^sub>D
K (G (F (H (D.dom (fst \<mu>\<nu>) \<star>\<^sub>D D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (F (\<Phi>\<^sub>H (D.dom (fst \<mu>\<nu>), D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
(K (G (F (H (D.dom (fst \<mu>\<nu>)) \<star>\<^sub>C H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (\<Phi>\<^sub>F (H (D.dom (fst \<mu>\<nu>)), H (D.dom (snd \<mu>\<nu>)))))) \<cdot>\<^sub>D
K (\<Phi>\<^sub>G (F (H (D.dom (fst \<mu>\<nu>))), F (H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
\<Phi>\<^sub>K (G (F (H (D.dom (fst \<mu>\<nu>)))), G (F (H (D.dom (snd \<mu>\<nu>)))))"
proof -
have "K (G (F (H (D.dom (fst \<mu>\<nu>)) \<star>\<^sub>C H (D.dom (snd \<mu>\<nu>))))) \<cdot>\<^sub>D
K (G (\<Phi>\<^sub>F (H (D.dom (fst \<mu>\<nu>)), H (D.dom (snd \<mu>\<nu>))))) =
K (G (\<Phi>\<^sub>F (H (D.dom (fst \<mu>\<nu>)), H (D.dom (snd \<mu>\<nu>)))))"
using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def
- B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
D.comp_cod_arr
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = KGFH.cmp \<mu>\<nu>"
using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def
- B.VV.arr_char B.VV.dom_simp B.VV.cod_simp
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
- D.VV.arr_char D.VV.dom_simp D.VV.cod_simp
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
+ D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_simp D.VV.cod_simp
F_.FF_def G_.FF_def H.FF_def K.FF_def
D.comp_assoc
by auto
finally show ?thesis by blast
qed
qed
ultimately show "equivalence_of_bicategories
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
FH.map FH.cmp KG.map KG.cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def
using B.bicategory_axioms D.bicategory_axioms FH.pseudofunctor_axioms
KG.pseudofunctor_axioms unit.pseudonatural_equivalence_axioms
counit.pseudonatural_equivalence_axioms I\<^sub>B.identity_pseudofunctor_axioms
I\<^sub>D.identity_pseudofunctor_axioms FH_KG.composite_pseudofunctor_axioms
KG_FH.composite_pseudofunctor_axioms
unfolding equivalence_of_bicategories_def by simp
qed
lemma is_equivalence_of_bicategories:
shows "equivalence_of_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
left_map left_cmp right_map right_cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
..
end
subsection "Equivalence with a Dense Sub-bicategory"
text \<open>
The purpose of this section is to show that, given a bicategory \<open>B\<close> and a sub-bicategory
defined by a ``dense'' set of objects of \<open>B\<close>, the embedding of the sub-bicategory in \<open>B\<close>
extends to an equivalence of bicategories. Here by ``dense'' we mean that every object
of \<open>B\<close> is equivalent to some object of the subbicategory.
\<close>
locale dense_subbicategory =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
subbicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B \<open>\<lambda>\<mu>. B.arr \<mu> \<and> src\<^sub>B \<mu> \<in> Obj \<and> trg\<^sub>B \<mu> \<in> Obj\<close>
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
and Obj :: "'b set" +
assumes dense: "\<And>a. B.obj a \<Longrightarrow> \<exists>a'. a' \<in> Obj \<and> B.equivalent_objects a' a"
begin
notation B.\<a>' ("\<a>\<^sub>B\<^sup>-\<^sup>1[_, _, _]")
notation B.lunit ("\<l>\<^sub>B[_]")
notation B.lunit' ("\<l>\<^sub>B\<^sup>-\<^sup>1[_]")
notation B.runit ("\<r>\<^sub>B[_]")
notation B.runit' ("\<r>\<^sub>B\<^sup>-\<^sup>1[_]")
notation comp (infixr "\<cdot>" 55)
notation hcomp (infixr "\<star>" 53)
notation in_hom ("\<guillemotleft>_ : _ \<Rightarrow> _\<guillemotright>")
notation in_hhom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation \<a> ("\<a>[_, _, _]")
notation \<a>' ("\<a>\<^sup>-\<^sup>1[_, _, _]")
notation lunit ("\<l>[_]")
notation lunit' ("\<l>\<^sup>-\<^sup>1[_]")
notation runit ("\<r>[_]")
notation runit' ("\<r>\<^sup>-\<^sup>1[_]")
abbreviation (input) Arr
where "Arr \<equiv> \<lambda>\<mu>. B.arr \<mu> \<and> src\<^sub>B \<mu> \<in> Obj \<and> trg\<^sub>B \<mu> \<in> Obj"
sublocale emb: embedding_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B Arr
..
abbreviation E
where "E \<equiv> emb.map"
abbreviation \<Phi>\<^sub>E
where "\<Phi>\<^sub>E \<equiv> emb.cmp"
text \<open>
We define a projection \<open>P\<close> by transporting arrows of \<open>B\<close> across chosen
equivalences between objects of \<open>B\<close> and objects of the sub-bicategory.
\<close>
definition P\<^sub>0
where "P\<^sub>0 a \<equiv> SOME a'. obj a' \<and> B.equivalent_objects a' a"
lemma P\<^sub>0_props:
assumes "B.obj a"
shows "obj (P\<^sub>0 a)"
and "B.equivalent_objects (P\<^sub>0 a) a"
and "B.equivalent_objects a a' \<Longrightarrow> P\<^sub>0 a = P\<^sub>0 a'"
and "P\<^sub>0 (P\<^sub>0 a) = P\<^sub>0 a"
and "B.obj (P\<^sub>0 a)"
and "P\<^sub>0 a \<in> Obj"
proof -
have "\<exists>a'. obj a' \<and> B.equivalent_objects a' a"
- using assms dense [of a] obj_char arr_char
+ using assms dense [of a] obj_char arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) B.equivalent_objects_def B.in_hhomE B.obj_def B.obj_simps(3))
hence 1: "obj (P\<^sub>0 a) \<and> B.equivalent_objects (P\<^sub>0 a) a"
unfolding P\<^sub>0_def
using someI_ex [of "\<lambda>a'. obj a' \<and> B.equivalent_objects a' a"] by blast
show "obj (P\<^sub>0 a)"
using 1 by simp
show 2: "B.equivalent_objects (P\<^sub>0 a) a"
using 1 by simp
show 3: "\<And>a'. B.equivalent_objects a a' \<Longrightarrow> P\<^sub>0 a = P\<^sub>0 a'"
unfolding P\<^sub>0_def
using B.equivalent_objects_symmetric B.equivalent_objects_transitive by meson
show "P\<^sub>0 (P\<^sub>0 a) = P\<^sub>0 a"
using 2 3 [of "P\<^sub>0 a"] B.equivalent_objects_symmetric by auto
show "B.obj (P\<^sub>0 a)"
using 1 B.equivalent_objects_def by auto
thus "P\<^sub>0 a \<in> Obj"
- using 1 3 obj_char arr_char by auto
+ using 1 3 obj_char arr_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
text \<open>
For each object \<open>a\<close> of \<open>B\<close>, we choose an adjoint equivalence from \<open>a\<close> to \<open>P\<^sub>0 a\<close>.
The use of adjoint equivalences is necessary in order to establish the required
coherence conditions.
\<close>
definition e
where "e a = (SOME e. \<guillemotleft>e : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright> \<and>
(\<exists>d \<eta> \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B e d \<eta> \<epsilon>))"
definition d
where "d a = (SOME d. \<exists>\<eta> \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) d \<eta> \<epsilon>)"
definition \<eta>
where "\<eta> a = (SOME \<eta>. \<exists>\<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) \<eta> \<epsilon>)"
definition \<epsilon>
where "\<epsilon> a = (SOME \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\<eta> a) \<epsilon>)"
lemma chosen_adjoint_equivalence:
assumes "B.obj a"
shows "adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\<eta> a) (\<epsilon> a)"
and "\<guillemotleft>e a : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>" and "B.ide (d a)" and "B.ide (e a)" and "B.iso (\<eta> a)" and "B.iso (\<epsilon> a)"
proof -
have "\<exists>e. \<guillemotleft>e : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright> \<and>
(\<exists>d \<eta> \<epsilon>. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B e d \<eta> \<epsilon>)"
proof -
obtain e where e: "\<guillemotleft>e : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright> \<and> B.equivalence_map e"
using assms P\<^sub>0_props(2) B.equivalent_objects_symmetric B.equivalent_objects_def
by meson
obtain d \<eta> \<epsilon> where d: "adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B e d \<eta> \<epsilon>"
using e B.equivalence_map_extends_to_adjoint_equivalence by blast
thus ?thesis
using e d by auto
qed
hence 1: "\<guillemotleft>e a : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright> \<and>
adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\<eta> a) (\<epsilon> a)"
- using d_def e_def \<eta>_def \<epsilon>_def arr_char
+ using d_def e_def \<eta>_def \<epsilon>_def arr_char\<^sub>S\<^sub>b\<^sub>C
someI_ex [of "\<lambda>e. \<guillemotleft>e : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright> \<and>
(\<exists>d \<eta> \<epsilon>. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B e d \<eta> \<epsilon>)"]
someI_ex [of "\<lambda>d. (\<exists>\<eta> \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) d \<eta> \<epsilon>)"]
someI_ex [of "\<lambda>\<eta>. \<exists>\<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\<eta> a) \<epsilon>"]
by simp
interpret adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B \<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using 1 by simp
show "adjoint_equivalence_in_bicategory
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\<eta> a) (\<epsilon> a)" ..
show "\<guillemotleft>e a : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>" using 1 by simp
show "B.ide (d a)" by simp
show "B.ide (e a)" by simp
show "B.iso (\<eta> a)" by simp
show "B.iso (\<epsilon> a)" by simp
qed
lemma equivalence_data_in_hom\<^sub>B [intro]:
assumes "B.obj a"
shows "\<guillemotleft>e a : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>" and "\<guillemotleft>d a : P\<^sub>0 a \<rightarrow>\<^sub>B a\<guillemotright>"
and "\<guillemotleft>e a : e a \<Rightarrow>\<^sub>B e a\<guillemotright>" and "\<guillemotleft>d a : d a \<Rightarrow>\<^sub>B d a\<guillemotright>"
and "\<guillemotleft>\<eta> a : a \<rightarrow>\<^sub>B a\<guillemotright>" and "\<guillemotleft>\<epsilon> a : P\<^sub>0 a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<eta> a : a \<Rightarrow>\<^sub>B d a \<star>\<^sub>B e a\<guillemotright>" and "\<guillemotleft>\<epsilon> a : e a \<star>\<^sub>B d a \<Rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>"
proof -
interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B \<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show e: "\<guillemotleft>e a : a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>"
using assms chosen_adjoint_equivalence by simp
show "\<guillemotleft>d a : P\<^sub>0 a \<rightarrow>\<^sub>B a\<guillemotright>" using e antipar by auto
show "\<guillemotleft>e a : e a \<Rightarrow>\<^sub>B e a\<guillemotright>" by auto
show "\<guillemotleft>d a : d a \<Rightarrow>\<^sub>B d a\<guillemotright>" by auto
show "\<guillemotleft>\<eta> a : a \<rightarrow>\<^sub>B a\<guillemotright>" using unit_in_hom e by auto
show "\<guillemotleft>\<epsilon> a : P\<^sub>0 a \<rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>" using counit_in_hom e by auto
show "\<guillemotleft>\<eta> a : a \<Rightarrow>\<^sub>B d a \<star>\<^sub>B e a\<guillemotright>" using unit_in_hom e by auto
show "\<guillemotleft>\<epsilon> a : e a \<star>\<^sub>B d a \<Rightarrow>\<^sub>B P\<^sub>0 a\<guillemotright>" using counit_in_hom e by auto
qed
lemma equivalence_data_simps\<^sub>B [simp]:
assumes "B.obj a"
shows "B.ide (d a)" and "B.ide (e a)" and "B.iso (\<eta> a)" and "B.iso (\<epsilon> a)"
and "src\<^sub>B (e a) = a" and "trg\<^sub>B (e a) = P\<^sub>0 a" and "src\<^sub>B (d a) = P\<^sub>0 a" and "trg\<^sub>B (d a) = a"
and "B.dom (e a) = e a" and "B.cod (e a) = e a"
and "B.dom (d a) = d a" and "B.cod (d a) = d a"
and "src\<^sub>B (\<eta> a) = a" and "trg\<^sub>B (\<eta> a) = a" and "src\<^sub>B (\<epsilon> a) = P\<^sub>0 a" and "trg\<^sub>B (\<epsilon> a) = P\<^sub>0 a"
and "B.dom (\<eta> a) = a" and "B.cod (\<eta> a) = d a \<star>\<^sub>B e a"
and "B.dom (\<epsilon> a) = e a \<star>\<^sub>B d a" and "B.cod (\<epsilon> a) = P\<^sub>0 a"
using assms chosen_adjoint_equivalence equivalence_data_in_hom\<^sub>B B.in_hhom_def
apply auto
by (meson B.in_homE)+
lemma equivalence_data_in_hom [intro]:
assumes "obj a"
shows "\<guillemotleft>e a : a \<rightarrow> P\<^sub>0 a\<guillemotright>" and "\<guillemotleft>d a : P\<^sub>0 a \<rightarrow> a\<guillemotright>"
and "\<guillemotleft>e a : e a \<Rightarrow> e a\<guillemotright>" and "\<guillemotleft>d a : d a \<Rightarrow> d a\<guillemotright>"
and "\<guillemotleft>\<eta> a : a \<rightarrow> a\<guillemotright>" and "\<guillemotleft>\<epsilon> a : P\<^sub>0 a \<rightarrow> P\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<eta> a : a \<Rightarrow> d a \<star> e a\<guillemotright>" and "\<guillemotleft>\<epsilon> a : e a \<star> d a \<Rightarrow> P\<^sub>0 a\<guillemotright>"
proof -
have a: "B.obj a \<and> a \<in> Obj"
- using assms obj_char arr_char by auto
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C by auto
have P\<^sub>0a: "obj (P\<^sub>0 a) \<and> P\<^sub>0 a \<in> Obj"
- using assms a P\<^sub>0_props [of a] arr_char obj_char by auto
+ using assms a P\<^sub>0_props [of a] arr_char\<^sub>S\<^sub>b\<^sub>C obj_char by auto
show ea: "\<guillemotleft>e a : a \<rightarrow> P\<^sub>0 a\<guillemotright>"
- using a P\<^sub>0a arr_char chosen_adjoint_equivalence(2) src_def trg_def by auto
+ using a P\<^sub>0a arr_char\<^sub>S\<^sub>b\<^sub>C chosen_adjoint_equivalence(2) src_def trg_def by auto
show "\<guillemotleft>e a : e a \<Rightarrow> e a\<guillemotright>"
- using a ea chosen_adjoint_equivalence(4) ide_char in_hom_char by auto
+ using a ea chosen_adjoint_equivalence(4) ide_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
show da: "\<guillemotleft>d a : P\<^sub>0 a \<rightarrow> a\<guillemotright>"
- using a P\<^sub>0a arr_char src_def trg_def P\<^sub>0_props(1) equivalence_data_in_hom\<^sub>B(2)
+ using a P\<^sub>0a arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def P\<^sub>0_props(1) equivalence_data_in_hom\<^sub>B(2)
by auto
show "\<guillemotleft>d a : d a \<Rightarrow> d a\<guillemotright>"
- using a da chosen_adjoint_equivalence(3) ide_char in_hom_char by auto
+ using a da chosen_adjoint_equivalence(3) ide_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
show \<eta>a: "\<guillemotleft>\<eta> a : a \<Rightarrow> d a \<star> e a\<guillemotright>"
proof
show 1: "arr (\<eta> a)"
- using assms a P\<^sub>0a da ea arr_char equivalence_data_in_hom\<^sub>B(7) hcomp_closed
+ using assms a P\<^sub>0a da ea arr_char\<^sub>S\<^sub>b\<^sub>C equivalence_data_in_hom\<^sub>B(7) hcomp_closed
by (metis (no_types, lifting) equivalence_data_in_hom\<^sub>B(5) B.in_hhom_def)
show "dom (\<eta> a) = a"
- using a 1 dom_char equivalence_data_in_hom\<^sub>B(7) by auto
+ using a 1 dom_char\<^sub>S\<^sub>b\<^sub>C equivalence_data_in_hom\<^sub>B(7) by auto
show "cod (\<eta> a) = d a \<star> e a"
- using a 1 da ea cod_char in_hom_char equivalence_data_in_hom\<^sub>B(7) hcomp_def
+ using a 1 da ea cod_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C equivalence_data_in_hom\<^sub>B(7) hcomp_def
by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE)
qed
show \<epsilon>a: "\<guillemotleft>\<epsilon> a : e a \<star> d a \<Rightarrow> P\<^sub>0 a\<guillemotright>"
proof
show 1: "arr (\<epsilon> a)"
- using assms a P\<^sub>0a da ea arr_char equivalence_data_in_hom\<^sub>B(8) hcomp_closed
+ using assms a P\<^sub>0a da ea arr_char\<^sub>S\<^sub>b\<^sub>C equivalence_data_in_hom\<^sub>B(8) hcomp_closed
by (metis (no_types, lifting) equivalence_data_in_hom\<^sub>B(6) B.in_hhom_def)
show "dom (\<epsilon> a) = e a \<star> d a"
- using a 1 da ea dom_char in_hom_char equivalence_data_in_hom\<^sub>B(8) hcomp_def
+ using a 1 da ea dom_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C equivalence_data_in_hom\<^sub>B(8) hcomp_def
by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE)
show "cod (\<epsilon> a) = P\<^sub>0 a"
- using a 1 cod_char equivalence_data_in_hom\<^sub>B(8) by auto
+ using a 1 cod_char\<^sub>S\<^sub>b\<^sub>C equivalence_data_in_hom\<^sub>B(8) by auto
qed
show "\<guillemotleft>\<eta> a : a \<rightarrow> a\<guillemotright>"
using assms \<eta>a src_def trg_def src_dom trg_dom
- by (metis (no_types, lifting) in_hhom_def in_hom_char obj_simps(2-3)
+ by (metis (no_types, lifting) in_hhom_def in_hom_char\<^sub>S\<^sub>b\<^sub>C obj_simps(2-3)
vconn_implies_hpar(1-2))
show "\<guillemotleft>\<epsilon> a : P\<^sub>0 a \<rightarrow> P\<^sub>0 a\<guillemotright>"
using P\<^sub>0a \<epsilon>a src_def trg_def src_cod trg_cod
- by (metis (no_types, lifting) in_hhom_def in_hom_char obj_simps(2-3)
+ by (metis (no_types, lifting) in_hhom_def in_hom_char\<^sub>S\<^sub>b\<^sub>C obj_simps(2-3)
vconn_implies_hpar(1-4))
qed
lemma equivalence_data_simps [simp]:
assumes "obj a"
shows "ide (d a)" and "ide (e a)" and "iso (\<eta> a)" and "iso (\<epsilon> a)"
and "src (e a) = a" and "trg (e a) = P\<^sub>0 a" and "src (d a) = P\<^sub>0 a" and "trg (d a) = a"
and "dom (e a) = e a" and "cod (e a) = e a"
and "dom (d a) = d a" and "cod (d a) = d a"
and "src (\<eta> a) = a" and "trg (\<eta> a) = a" and "src (\<epsilon> a) = P\<^sub>0 a" and "trg (\<epsilon> a) = P\<^sub>0 a"
and "dom (\<eta> a) = a" and "cod (\<eta> a) = d a \<star>\<^sub>B e a"
and "dom (\<epsilon> a) = e a \<star>\<^sub>B d a" and "cod (\<epsilon> a) = P\<^sub>0 a"
(* These can be done in a one-liner, but it takes 20 sec. *)
- using assms equivalence_data_in_hom(2) ide_char obj_char apply auto[1]
- using assms equivalence_data_in_hom(1) ide_char obj_char apply auto[1]
- using assms obj_char arr_char iso_char B.iso_is_arr apply auto[1]
- using assms obj_char arr_char iso_char P\<^sub>0_props B.iso_is_arr apply auto[1]
- using assms obj_char arr_char P\<^sub>0_props src_def apply auto[1]
- using assms obj_char arr_char P\<^sub>0_props trg_def apply auto[1]
- using assms obj_char arr_char P\<^sub>0_props src_def apply auto[1]
- using assms obj_char arr_char P\<^sub>0_props trg_def apply auto[1]
- using assms obj_char arr_char dom_char P\<^sub>0_props apply auto[1]
- using assms obj_char arr_char cod_char P\<^sub>0_props apply auto[1]
- using assms obj_char arr_char dom_char P\<^sub>0_props apply auto[1]
- using assms obj_char arr_char cod_char P\<^sub>0_props apply auto[1]
- using assms obj_char arr_char B.iso_is_arr src_def apply auto[1]
- using assms obj_char arr_char B.iso_is_arr trg_def apply auto[1]
- using assms obj_char arr_char dom_char P\<^sub>0_props B.iso_is_arr src_def apply auto[1]
- using assms obj_char arr_char cod_char P\<^sub>0_props B.iso_is_arr trg_def apply auto[1]
- using assms obj_char arr_char dom_char B.iso_is_arr apply auto[1]
- using assms obj_char arr_char cod_char B.iso_is_arr apply auto[1]
- using assms obj_char arr_char dom_char P\<^sub>0_props B.iso_is_arr apply auto[1]
- using assms obj_char arr_char cod_char P\<^sub>0_props B.iso_is_arr by auto[1]
+ using assms equivalence_data_in_hom(2) ide_char\<^sub>S\<^sub>b\<^sub>C obj_char apply auto[1]
+ using assms equivalence_data_in_hom(1) ide_char\<^sub>S\<^sub>b\<^sub>C obj_char apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C iso_char\<^sub>S\<^sub>b\<^sub>C B.iso_is_arr apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C iso_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props B.iso_is_arr apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props src_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props trg_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props src_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props trg_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C B.iso_is_arr src_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C B.iso_is_arr trg_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props B.iso_is_arr src_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props B.iso_is_arr trg_def apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C B.iso_is_arr apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C B.iso_is_arr apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props B.iso_is_arr apply auto[1]
+ using assms obj_char arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props B.iso_is_arr by auto[1]
definition P
where "P \<mu> = e (trg\<^sub>B \<mu>) \<star>\<^sub>B \<mu> \<star>\<^sub>B d (src\<^sub>B \<mu>)"
lemma P_in_hom\<^sub>B [intro]:
assumes "B.arr \<mu>"
shows "\<guillemotleft>P \<mu> : P\<^sub>0 (src\<^sub>B \<mu>) \<rightarrow>\<^sub>B P\<^sub>0 (trg\<^sub>B \<mu>)\<guillemotright>"
and "\<guillemotleft>P \<mu> : P (B.dom \<mu>) \<Rightarrow>\<^sub>B P (B.cod \<mu>)\<guillemotright>"
unfolding P_def using assms by auto
lemma P_simps\<^sub>B [simp]:
assumes "B.arr \<mu>"
shows "B.arr (P \<mu>)"
and "src\<^sub>B (P \<mu>) = P\<^sub>0 (src\<^sub>B \<mu>)" and "trg\<^sub>B (P \<mu>) = P\<^sub>0 (trg\<^sub>B \<mu>)"
and "B.dom (P \<mu>) = P (B.dom \<mu>)" and "B.cod (P \<mu>) = P (B.cod \<mu>)"
using assms P_in_hom\<^sub>B by blast+
lemma P_in_hom [intro]:
assumes "B.arr \<mu>"
shows "\<guillemotleft>P \<mu> : P\<^sub>0 (src\<^sub>B \<mu>) \<rightarrow> P\<^sub>0 (trg\<^sub>B \<mu>)\<guillemotright>"
and "\<guillemotleft>P \<mu> : P (B.dom \<mu>) \<Rightarrow> P (B.cod \<mu>)\<guillemotright>"
proof -
show 1: "\<guillemotleft>P \<mu> : P\<^sub>0 (src\<^sub>B \<mu>) \<rightarrow> P\<^sub>0 (trg\<^sub>B \<mu>)\<guillemotright>"
- using assms P_in_hom\<^sub>B(1) arr_char src_def trg_def P\<^sub>0_props(1)
+ using assms P_in_hom\<^sub>B(1) arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def P\<^sub>0_props(1)
by (metis (no_types, lifting) in_hhom_def obj_char B.in_hhomE B.obj_simps(2)
B.obj_src B.obj_trg)
show "\<guillemotleft>P \<mu> : P (B.dom \<mu>) \<Rightarrow> P (B.cod \<mu>)\<guillemotright>"
- using assms 1 dom_char cod_char
+ using assms 1 dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C
by (intro in_homI) auto
qed
lemma P_simps [simp]:
assumes "B.arr \<mu>"
shows "arr (P \<mu>)"
and "src (P \<mu>) = P\<^sub>0 (src\<^sub>B \<mu>)" and "trg (P \<mu>) = P\<^sub>0 (trg\<^sub>B \<mu>)"
and "dom (P \<mu>) = P (B.dom \<mu>)" and "cod (P \<mu>) = P (B.cod \<mu>)"
using assms P_in_hom by blast+
interpretation P: "functor" V\<^sub>B comp P
proof
show "\<And>\<mu>. \<not> B.arr \<mu> \<Longrightarrow> P \<mu> = null"
using P_def null_char B.hseq_char B.hseq_char' by auto
have 0: "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> B.arr (P \<mu>)"
by simp
show 1: "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> arr (P \<mu>)"
using P_simps\<^sub>B(1) by simp
show 2: "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> dom (P \<mu>) = P (B.dom \<mu>)"
using 1 dom_simp by auto
show 3: "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> cod (P \<mu>) = P (B.cod \<mu>)"
using 1 cod_simp by auto
show "\<And>\<mu> \<nu>. B.seq \<mu> \<nu> \<Longrightarrow> P (\<mu> \<cdot>\<^sub>B \<nu>) = P \<mu> \<cdot> P \<nu>"
proof -
fix \<mu> \<nu>
assume seq: "B.seq \<mu> \<nu>"
show "P (\<mu> \<cdot>\<^sub>B \<nu>) = P \<mu> \<cdot> P \<nu>"
proof -
have 4: "P (\<mu> \<cdot>\<^sub>B \<nu>) = e (trg\<^sub>B \<mu>) \<star>\<^sub>B \<mu> \<cdot>\<^sub>B \<nu> \<star>\<^sub>B d (src\<^sub>B \<mu>)"
unfolding P_def using seq B.src_vcomp B.trg_vcomp by simp
also have "... = P \<mu> \<cdot>\<^sub>B P \<nu>"
unfolding P_def
using seq 0 4 B.whisker_left B.whisker_right
by (metis equivalence_data_simps\<^sub>B(1-2) B.hseqE B.obj_trg B.seqE B.src_vcomp
B.vseq_implies_hpar(2))
also have "... = P \<mu> \<cdot> P \<nu>"
- using seq 1 seq_char comp_char by auto
+ using seq 1 seq_char\<^sub>S\<^sub>b\<^sub>C comp_char by auto
finally show ?thesis by blast
qed
qed
qed
interpretation faithful_functor V\<^sub>B comp P
proof
fix \<mu> \<mu>'
assume par: "B.par \<mu> \<mu>'"
have 1: "src\<^sub>B \<mu> = src\<^sub>B \<mu>' \<and> trg\<^sub>B \<mu> = trg\<^sub>B \<mu>'"
using par by (metis B.src_dom B.trg_dom)
assume eq: "P \<mu> = P \<mu>'"
interpret src: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (src\<^sub>B \<mu>)\<close> \<open>d (src\<^sub>B \<mu>)\<close> \<open>\<eta> (src\<^sub>B \<mu>)\<close> \<open>\<epsilon> (src\<^sub>B \<mu>)\<close>
using par chosen_adjoint_equivalence by simp
interpret trg: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (trg\<^sub>B \<mu>)\<close> \<open>d (trg\<^sub>B \<mu>)\<close> \<open>\<eta> (trg\<^sub>B \<mu>)\<close> \<open>\<epsilon> (trg\<^sub>B \<mu>)\<close>
using par chosen_adjoint_equivalence by simp
show "\<mu> = \<mu>'"
using eq par 1
unfolding P_def
using B.equivalence_cancel_left [of "e (trg\<^sub>B \<mu>)" "\<mu> \<star>\<^sub>B d (src\<^sub>B \<mu>)" "\<mu>' \<star>\<^sub>B d (src\<^sub>B \<mu>')"]
B.equivalence_cancel_right [of "d (src\<^sub>B \<mu>)" \<mu> \<mu>']
B.equivalence_map_def src.dual_equivalence trg.equivalence_in_bicategory_axioms
by auto
qed
interpretation P: weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B comp src trg P
proof
fix \<mu>
assume \<mu>: "B.arr \<mu>"
interpret src: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (src\<^sub>B \<mu>)\<close> \<open>d (src\<^sub>B \<mu>)\<close> \<open>\<eta> (src\<^sub>B \<mu>)\<close> \<open>\<epsilon> (src\<^sub>B \<mu>)\<close>
using \<mu> chosen_adjoint_equivalence by simp
show "isomorphic (P (src\<^sub>B \<mu>)) (src (P \<mu>))"
proof (unfold isomorphic_def)
show "\<exists>f. \<guillemotleft>f : P (src\<^sub>B \<mu>) \<Rightarrow> src (P \<mu>)\<guillemotright> \<and> iso f"
proof -
let ?\<phi> = "\<epsilon> (src\<^sub>B \<mu>) \<cdot>\<^sub>B (e (src\<^sub>B \<mu>) \<star>\<^sub>B \<l>\<^sub>B[d (src\<^sub>B \<mu>)])"
have "\<guillemotleft>?\<phi> : P (src\<^sub>B \<mu>) \<Rightarrow> src (P \<mu>)\<guillemotright> \<and> iso ?\<phi>"
proof -
have 1: "\<guillemotleft>?\<phi> : P (src\<^sub>B \<mu>) \<Rightarrow>\<^sub>B P\<^sub>0 (src\<^sub>B \<mu>)\<guillemotright> \<and> B.iso ?\<phi>"
unfolding P_def using \<mu> by auto
moreover have "arr ?\<phi> \<and> arr (B.inv ?\<phi>)"
- using \<mu> 1 arr_char P\<^sub>0_props(1)
+ using \<mu> 1 arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1)
by (metis (no_types, lifting) P\<^sub>0_props(6) B.arrI B.arr_inv equivalence_data_simps\<^sub>B(16)
B.obj_src src.counit_simps(4-5) B.src_inv B.src_vcomp
B.trg_inv B.trg_vcomp)
moreover have "P\<^sub>0 (src\<^sub>B \<mu>) = src (P \<mu>)"
using \<mu> by simp
ultimately show ?thesis
- using \<mu> in_hom_char arr_char iso_char
+ using \<mu> in_hom_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C iso_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) src_closed equivalence_data_simps\<^sub>B(6)
B.obj_src P.preserves_arr src.counit_simps(4) B.src.preserves_arr B.src_vcomp)
qed
thus ?thesis by blast
qed
qed
interpret trg: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (trg\<^sub>B \<mu>)\<close> \<open>d (trg\<^sub>B \<mu>)\<close> \<open>\<eta> (trg\<^sub>B \<mu>)\<close> \<open>\<epsilon> (trg\<^sub>B \<mu>)\<close>
using \<mu> chosen_adjoint_equivalence by simp
show "isomorphic (P (trg\<^sub>B \<mu>)) (trg (P \<mu>))"
proof (unfold isomorphic_def)
show "\<exists>f. \<guillemotleft>f : P (trg\<^sub>B \<mu>) \<Rightarrow> trg (P \<mu>)\<guillemotright> \<and> iso f"
proof -
let ?\<psi> = "\<epsilon> (trg\<^sub>B \<mu>) \<cdot>\<^sub>B (e (trg\<^sub>B \<mu>) \<star>\<^sub>B \<l>\<^sub>B[d (trg\<^sub>B \<mu>)])"
have "\<guillemotleft>?\<psi> : P (trg\<^sub>B \<mu>) \<Rightarrow> trg (P \<mu>)\<guillemotright> \<and> iso ?\<psi>"
proof -
have 1: "\<guillemotleft>?\<psi> : P (trg\<^sub>B \<mu>) \<Rightarrow>\<^sub>B P\<^sub>0 (trg\<^sub>B \<mu>)\<guillemotright> \<and> B.iso ?\<psi>"
unfolding P_def using \<mu> by auto
moreover have "arr ?\<psi> \<and> arr (B.inv ?\<psi>)"
- using \<mu> 1 arr_char P\<^sub>0_props(1)
+ using \<mu> 1 arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1)
by (metis (no_types, lifting) obj_char B.arrI B.arr_inv B.obj_trg B.src_inv B.trg_inv
B.vconn_implies_hpar(1-4))
moreover have "arr (P (trg\<^sub>B \<mu>)) \<and> arr (P\<^sub>0 (trg\<^sub>B \<mu>))"
- using \<mu> arr_char P\<^sub>0_props(1) P_simps(1) obj_char by auto
+ using \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) P_simps(1) obj_char by auto
moreover have "P\<^sub>0 (trg\<^sub>B \<mu>) = trg (P \<mu>)"
using \<mu> by simp
ultimately show ?thesis
- using in_hom_char iso_char by force
+ using in_hom_char\<^sub>S\<^sub>b\<^sub>C iso_char\<^sub>S\<^sub>b\<^sub>C by force
qed
thus ?thesis by blast
qed
qed
qed
text \<open>
The following seems to be needed to avoid non-confluent simplifications,
\emph{e.g.}~of \<open>S.src (P \<mu>)\<close> to \<open>P.map\<^sub>0 a\<close> and to \<open>P\<^sub>0 a\<close>.
\<close>
lemma P_map\<^sub>0_simp [simp]:
assumes "B.obj a"
shows "P.map\<^sub>0 a = P\<^sub>0 a"
using assms P.map\<^sub>0_def B.obj_simps(1-2) by simp
interpretation HoPP: composite_functor B.VV.comp VV.comp comp
P.FF \<open>\<lambda>\<mu>\<nu>. hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
..
interpretation PoH: composite_functor B.VV.comp V\<^sub>B comp \<open>(\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)\<close> P
..
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
definition CMP
where "CMP f g \<equiv> (e (trg\<^sub>B f) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d (src\<^sub>B g)]) \<cdot>\<^sub>B
(e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B
\<l>\<^sub>B[g \<star>\<^sub>B d (src\<^sub>B g)] \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B g)) \<star>\<^sub>B g \<star>\<^sub>B d (src\<^sub>B g))) \<cdot>\<^sub>B
(e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d (src\<^sub>B f), e (trg\<^sub>B g), g \<star>\<^sub>B d (src\<^sub>B g)]) \<cdot>\<^sub>B
\<a>\<^sub>B[e (trg\<^sub>B f), f, d (src\<^sub>B f) \<star>\<^sub>B P g] \<cdot>\<^sub>B
\<a>\<^sub>B[e (trg\<^sub>B f) \<star>\<^sub>B f, d (src\<^sub>B f), P g] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e (trg\<^sub>B f), f, d (src\<^sub>B f)] \<star>\<^sub>B P g)"
text \<open>
The 2-cell \<open>CMP f g\<close> has the right type to be a compositor for a pseudofunctor
whose underlying mapping is \<open>P\<close>.
\<close>
lemma CMP_in_hom [intro]:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "\<guillemotleft>CMP f g : P\<^sub>0 (src\<^sub>B g) \<rightarrow> P\<^sub>0 (trg\<^sub>B f)\<guillemotright>"
and "\<guillemotleft>CMP f g : P f \<star> P g \<Rightarrow> P (f \<star>\<^sub>B g)\<guillemotright>"
and "\<guillemotleft>CMP f g : P\<^sub>0 (src\<^sub>B g) \<rightarrow>\<^sub>B P\<^sub>0 (trg\<^sub>B f)\<guillemotright>"
and "\<guillemotleft>CMP f g : P f \<star>\<^sub>B P g \<Rightarrow>\<^sub>B P (f \<star>\<^sub>B g)\<guillemotright>"
proof -
show 1: "\<guillemotleft>CMP f g : P f \<star>\<^sub>B P g \<Rightarrow>\<^sub>B P (f \<star>\<^sub>B g)\<guillemotright>"
using assms
by (unfold P_def CMP_def, intro B.comp_in_homI' B.seqI B.hseqI') auto
show 2: "\<guillemotleft>CMP f g : P\<^sub>0 (src\<^sub>B g) \<rightarrow>\<^sub>B P\<^sub>0 (trg\<^sub>B f)\<guillemotright>"
using assms 1 B.src_cod B.trg_cod B.vconn_implies_hpar(1-2) by auto
show 3: "\<guillemotleft>CMP f g : P f \<star> P g \<Rightarrow> P (f \<star>\<^sub>B g)\<guillemotright>"
- using assms 1 B.arrI B.vconn_implies_hpar(1-2) P_simps(1) arr_char hcomp_eqI
- hseqI' in_hom_char
+ using assms 1 B.arrI B.vconn_implies_hpar(1-2) P_simps(1) arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_eqI
+ hseqI' in_hom_char\<^sub>S\<^sub>b\<^sub>C
by force
show "\<guillemotleft>CMP f g : P\<^sub>0 (src\<^sub>B g) \<rightarrow> P\<^sub>0 (trg\<^sub>B f)\<guillemotright>"
- using 2 3 arr_char src_def trg_def by fastforce
+ using 2 3 arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def by fastforce
qed
lemma CMP_simps [simp]:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "arr (CMP f g)"
and "src (CMP f g) = P\<^sub>0 (src\<^sub>B g)" and "trg (CMP f g) = P\<^sub>0 (trg\<^sub>B f)"
and "dom (CMP f g) = P f \<star> P g" and "cod (CMP f g) = P (f \<star>\<^sub>B g)"
and "B.arr (CMP f g)"
and "src\<^sub>B (CMP f g) = P\<^sub>0 (src\<^sub>B g)" and "trg\<^sub>B (CMP f g) = P\<^sub>0 (trg\<^sub>B f)"
and "B.dom (CMP f g) = P f \<star>\<^sub>B P g" and "B.cod (CMP f g) = P (f \<star>\<^sub>B g)"
using assms CMP_in_hom [of f g] by auto
lemma iso_CMP:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "iso (CMP f g)"
and "B.iso (CMP f g)"
proof -
show "B.iso (CMP f g)"
unfolding CMP_def P_def
- using assms B.VV.arr_char P.as_nat_iso.components_are_iso
+ using assms B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P.as_nat_iso.components_are_iso
by (intro B.isos_compose B.iso_hcomp) auto
thus "iso (CMP f g)"
- using assms iso_char arr_char CMP_simps(1) by auto
+ using assms iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C CMP_simps(1) by auto
qed
abbreviation (input) SRC
where "SRC \<mu> \<equiv> d (src\<^sub>B \<mu>) \<star>\<^sub>B e (src\<^sub>B \<mu>)"
abbreviation (input) TRG
where "TRG \<mu> \<equiv> d (trg\<^sub>B \<mu>) \<star>\<^sub>B e (trg\<^sub>B \<mu>)"
definition LUNIT
where "LUNIT f \<equiv> \<l>\<^sub>B[f] \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B f)) \<star>\<^sub>B f)"
definition RUNIT
where "RUNIT f \<equiv> \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B f)))"
text \<open>
Here we prove a series of results that would be automatic if we had some notion of
``bicategory with \<open>SRC\<close> and \<open>TRG\<close> as alternative source and target''.
Perhaps this idea can be developed in future work and used to simplify the overall
development.
\<close>
lemma LUNIT_in_hom [intro]:
assumes "B.ide f"
shows "\<guillemotleft>LUNIT f : src\<^sub>B f \<rightarrow>\<^sub>B trg\<^sub>B f\<guillemotright>"
and "\<guillemotleft>LUNIT f : TRG f \<star>\<^sub>B f \<Rightarrow>\<^sub>B f\<guillemotright>"
proof -
interpret e_trg_f: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (trg\<^sub>B f)\<close> \<open>d (trg\<^sub>B f)\<close> \<open>\<eta> (trg\<^sub>B f)\<close> \<open>\<epsilon> (trg\<^sub>B f)\<close>
using assms chosen_adjoint_equivalence by simp
show "\<guillemotleft>LUNIT f : src\<^sub>B f \<rightarrow>\<^sub>B trg\<^sub>B f\<guillemotright>"
unfolding LUNIT_def
using assms e_trg_f.unit_is_iso by auto
show "\<guillemotleft>LUNIT f : TRG f \<star>\<^sub>B f \<Rightarrow>\<^sub>B f\<guillemotright>"
unfolding LUNIT_def
using assms e_trg_f.unit_is_iso by auto
qed
lemma LUNIT_simps [simp]:
assumes "B.ide f"
shows "B.arr (LUNIT f)"
and "src\<^sub>B (LUNIT f) = src\<^sub>B f" and "trg\<^sub>B (LUNIT f) = trg\<^sub>B f"
and "B.dom (LUNIT f) = TRG f \<star>\<^sub>B f"
and "B.cod (LUNIT f) = f"
using assms LUNIT_in_hom by auto
lemma RUNIT_in_hom [intro]:
assumes "B.ide f"
shows "\<guillemotleft>RUNIT f : src\<^sub>B f \<rightarrow>\<^sub>B trg\<^sub>B f\<guillemotright>"
and "\<guillemotleft>RUNIT f : f \<star>\<^sub>B SRC f \<Rightarrow>\<^sub>B f\<guillemotright>"
proof -
interpret e_src_f: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (src\<^sub>B f)\<close> \<open>d (src\<^sub>B f)\<close> \<open>\<eta> (src\<^sub>B f)\<close> \<open>\<epsilon> (src\<^sub>B f)\<close>
using assms chosen_adjoint_equivalence by simp
show "\<guillemotleft>RUNIT f : src\<^sub>B f \<rightarrow>\<^sub>B trg\<^sub>B f\<guillemotright>"
unfolding RUNIT_def
using assms e_src_f.unit_is_iso by auto
show "\<guillemotleft>RUNIT f : f \<star>\<^sub>B SRC f \<Rightarrow>\<^sub>B f\<guillemotright>"
unfolding RUNIT_def
using assms e_src_f.unit_is_iso by auto
qed
lemma RUNIT_simps [simp]:
assumes "B.ide f"
shows "B.arr (RUNIT f)"
and "src\<^sub>B (RUNIT f) = src\<^sub>B f" and "trg\<^sub>B (RUNIT f) = trg\<^sub>B f"
and "B.dom (RUNIT f) = f \<star>\<^sub>B SRC f"
and "B.cod (RUNIT f) = f"
using assms RUNIT_in_hom by auto
lemma iso_LUNIT:
assumes "B.ide f"
shows "B.iso (LUNIT f)"
by (simp add: assms B.isos_compose LUNIT_def)
lemma iso_RUNIT:
assumes "B.ide f"
shows "B.iso (RUNIT f)"
by (simp add: assms B.isos_compose RUNIT_def)
lemma LUNIT_naturality:
assumes "B.arr \<mu>"
shows "\<mu> \<cdot>\<^sub>B LUNIT (B.dom \<mu>) = LUNIT (B.cod \<mu>) \<cdot>\<^sub>B (TRG \<mu> \<star>\<^sub>B \<mu>)"
proof -
interpret e_trg_\<mu>: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (trg\<^sub>B \<mu>)\<close> \<open>d (trg\<^sub>B \<mu>)\<close> \<open>\<eta> (trg\<^sub>B \<mu>)\<close> \<open>\<epsilon> (trg\<^sub>B \<mu>)\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis
proof -
have "\<mu> \<cdot>\<^sub>B LUNIT (B.dom \<mu>) = (\<mu> \<cdot>\<^sub>B \<l>\<^sub>B[B.dom \<mu>]) \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B \<mu>)) \<star>\<^sub>B B.dom \<mu>)"
unfolding LUNIT_def
using assms B.comp_assoc by simp
also have "... = \<l>\<^sub>B[B.cod \<mu>] \<cdot>\<^sub>B (trg\<^sub>B \<mu> \<star>\<^sub>B \<mu>) \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B \<mu>)) \<star>\<^sub>B B.dom \<mu>)"
using assms B.lunit_naturality B.comp_assoc by simp
also have "... = \<l>\<^sub>B[B.cod \<mu>] \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B \<mu>)) \<star>\<^sub>B \<mu>)"
using assms B.interchange [of "trg\<^sub>B \<mu>" "B.inv (\<eta> (trg\<^sub>B \<mu>))" \<mu> "B.dom \<mu>"]
e_trg_\<mu>.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
also have "... = \<l>\<^sub>B[B.cod \<mu>] \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B \<mu>)) \<star>\<^sub>B B.cod \<mu>) \<cdot>\<^sub>B (TRG \<mu> \<star>\<^sub>B \<mu>)"
using assms B.interchange [of "B.inv (\<eta> (trg\<^sub>B \<mu>))" "TRG \<mu>" "B.cod \<mu>" \<mu>]
e_trg_\<mu>.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
also have "... = LUNIT (B.cod \<mu>) \<cdot>\<^sub>B (TRG \<mu> \<star>\<^sub>B \<mu>)"
unfolding LUNIT_def
using assms B.comp_assoc by simp
finally show ?thesis by simp
qed
qed
lemma RUNIT_naturality:
assumes "B.arr \<mu>"
shows "\<mu> \<cdot>\<^sub>B RUNIT (B.dom \<mu>) = RUNIT (B.cod \<mu>) \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B SRC \<mu>)"
proof -
interpret e_src_\<mu>: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (src\<^sub>B \<mu>)\<close> \<open>d (src\<^sub>B \<mu>)\<close> \<open>\<eta> (src\<^sub>B \<mu>)\<close> \<open>\<epsilon> (src\<^sub>B \<mu>)\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis
proof -
have "\<mu> \<cdot>\<^sub>B RUNIT (B.dom \<mu>) =
(\<mu> \<cdot>\<^sub>B \<r>\<^sub>B[B.dom \<mu>]) \<cdot>\<^sub>B (B.dom \<mu> \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B \<mu>)))"
unfolding RUNIT_def
using assms B.comp_assoc by simp
also have "... = \<r>\<^sub>B[B.cod \<mu>] \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B src\<^sub>B \<mu>) \<cdot>\<^sub>B (B.dom \<mu> \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B \<mu>)))"
using assms B.runit_naturality B.comp_assoc by simp
also have "... = \<r>\<^sub>B[B.cod \<mu>] \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B \<mu>)))"
using assms B.interchange [of \<mu> "B.dom \<mu>" "src\<^sub>B \<mu>" "B.inv (\<eta> (src\<^sub>B \<mu>))"]
e_src_\<mu>.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
also have "... = \<r>\<^sub>B[B.cod \<mu>] \<cdot>\<^sub>B (B.cod \<mu> \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B \<mu>))) \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B SRC \<mu>)"
using assms B.interchange [of "B.cod \<mu>" \<mu> "B.inv (\<eta> (src\<^sub>B \<mu>))" "SRC \<mu>"]
e_src_\<mu>.unit_is_iso B.comp_arr_dom B.comp_cod_arr
by simp
also have "... = RUNIT (B.cod \<mu>) \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B SRC \<mu>)"
unfolding RUNIT_def
using assms B.comp_assoc by simp
finally show ?thesis by simp
qed
qed
lemma LUNIT_hcomp:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "LUNIT (f \<star>\<^sub>B g) \<cdot>\<^sub>B \<a>\<^sub>B[d (trg\<^sub>B f) \<star>\<^sub>B e (trg\<^sub>B f), f, g] = LUNIT f \<star>\<^sub>B g"
proof -
interpret e_trg_f: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (trg\<^sub>B f)\<close> \<open>d (trg\<^sub>B f)\<close> \<open>\<eta> (trg\<^sub>B f)\<close> \<open>\<epsilon> (trg\<^sub>B f)\<close>
using assms chosen_adjoint_equivalence by simp
have "LUNIT (f \<star>\<^sub>B g) \<cdot>\<^sub>B \<a>\<^sub>B[TRG f, f, g] =
\<l>\<^sub>B[f \<star>\<^sub>B g] \<cdot>\<^sub>B (B.inv (\<eta> (trg\<^sub>B f)) \<star>\<^sub>B f \<star>\<^sub>B g) \<cdot>\<^sub>B \<a>\<^sub>B[TRG f, f, g]"
unfolding LUNIT_def
using assms B.comp_assoc by simp
also have "... = (\<l>\<^sub>B[f \<star>\<^sub>B g] \<cdot>\<^sub>B \<a>\<^sub>B[trg\<^sub>B f, f, g]) \<cdot>\<^sub>B ((B.inv (\<eta> (trg\<^sub>B f)) \<star>\<^sub>B f) \<star>\<^sub>B g)"
using assms B.assoc_naturality [of "B.inv (\<eta> (trg\<^sub>B f))" f g] e_trg_f.unit_is_iso
B.comp_assoc
by simp
also have "... = (\<l>\<^sub>B[f] \<star>\<^sub>B g) \<cdot>\<^sub>B ((B.inv (\<eta> (trg\<^sub>B f)) \<star>\<^sub>B f) \<star>\<^sub>B g)"
using assms B.lunit_hcomp by simp
also have "... = LUNIT f \<star>\<^sub>B g"
using assms LUNIT_def LUNIT_simps(1) B.whisker_right by auto
finally show ?thesis by simp
qed
lemma RUNIT_hcomp:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "RUNIT (f \<star>\<^sub>B g) = (f \<star>\<^sub>B RUNIT g) \<cdot>\<^sub>B \<a>\<^sub>B[f, g, SRC g]"
proof -
interpret e_src_g: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (src\<^sub>B g)\<close> \<open>d (src\<^sub>B g)\<close> \<open>\<eta> (src\<^sub>B g)\<close> \<open>\<epsilon> (src\<^sub>B g)\<close>
using assms chosen_adjoint_equivalence by simp
have "(f \<star>\<^sub>B RUNIT g) \<cdot>\<^sub>B \<a>\<^sub>B[f, g, SRC g] =
(f \<star>\<^sub>B \<r>\<^sub>B[g]) \<cdot>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B g))) \<cdot>\<^sub>B \<a>\<^sub>B[f, g, SRC g]"
unfolding RUNIT_def
using assms B.whisker_left e_src_g.unit_is_iso B.comp_assoc by simp
also have "... = ((f \<star>\<^sub>B \<r>\<^sub>B[g]) \<cdot>\<^sub>B \<a>\<^sub>B[f, g, src\<^sub>B g]) \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B g)))"
using assms B.assoc_naturality [of f g "B.inv (\<eta> (src\<^sub>B g))"] e_src_g.unit_is_iso
B.comp_assoc
by simp
also have "... = \<r>\<^sub>B[f \<star>\<^sub>B g] \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B g)))"
using assms B.runit_hcomp by simp
also have "... = RUNIT (f \<star>\<^sub>B g)"
using assms RUNIT_def by simp
finally show ?thesis by simp
qed
lemma TRIANGLE:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "(f \<star>\<^sub>B LUNIT g) \<cdot>\<^sub>B \<a>\<^sub>B[f, SRC f, g] = RUNIT f \<star>\<^sub>B g"
proof -
interpret e_trg_g: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e (trg\<^sub>B g)\<close> \<open>d (trg\<^sub>B g)\<close> \<open>\<eta> (trg\<^sub>B g)\<close> \<open>\<epsilon> (trg\<^sub>B g)\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis
proof -
have "(f \<star>\<^sub>B LUNIT g) \<cdot>\<^sub>B \<a>\<^sub>B[f, SRC f, g] =
(f \<star>\<^sub>B \<l>\<^sub>B[g]) \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> (trg\<^sub>B g)) \<star>\<^sub>B g) \<cdot>\<^sub>B \<a>\<^sub>B[f, SRC f, g]"
using assms B.whisker_left e_trg_g.unit_is_iso LUNIT_def LUNIT_simps(1) B.comp_assoc
by auto
also have "... = ((f \<star>\<^sub>B \<l>\<^sub>B[g]) \<cdot>\<^sub>B \<a>\<^sub>B[f, src\<^sub>B f, g]) \<cdot>\<^sub>B ((f \<star>\<^sub>B B.inv (\<eta> (trg\<^sub>B g))) \<star>\<^sub>B g)"
using assms B.assoc_naturality [of f "B.inv (\<eta> (trg\<^sub>B g))" g] e_trg_g.unit_is_iso
B.comp_assoc
by auto
also have "... = (\<r>\<^sub>B[f] \<star>\<^sub>B g) \<cdot>\<^sub>B ((f \<star>\<^sub>B B.inv (\<eta> (trg\<^sub>B g))) \<star>\<^sub>B g)"
using assms B.triangle by simp
also have "... = RUNIT f \<star>\<^sub>B g"
using assms B.whisker_right e_trg_g.unit_is_iso RUNIT_def RUNIT_simps
by metis
finally show ?thesis by simp
qed
qed
text \<open>
The \<open>CMP f g\<close> also satisfy the naturality conditions required of compositors.
\<close>
lemma CMP_naturality:
assumes "B.arr \<mu>" and "B.arr \<nu>" and "src\<^sub>B \<mu> = trg\<^sub>B \<nu>"
shows "CMP (B.cod \<mu>) (B.cod \<nu>) \<cdot>\<^sub>B (P \<mu> \<star>\<^sub>B P \<nu>)
= P (\<mu> \<star>\<^sub>B \<nu>) \<cdot>\<^sub>B CMP (B.dom \<mu>) (B.dom \<nu>)"
proof -
let ?a = "src\<^sub>B \<nu>"
let ?b = "src\<^sub>B \<mu>"
let ?c = "trg\<^sub>B \<mu>"
let ?f = "B.dom \<mu>"
let ?g = "B.cod \<mu>"
let ?h = "B.dom \<nu>"
let ?k = "B.cod \<nu>"
have "CMP ?g ?k \<cdot>\<^sub>B (P \<mu> \<star>\<^sub>B P \<nu>)
= (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, ?g, d ?b \<star>\<^sub>B P ?k] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B ?g, d ?b, P ?k] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \<star>\<^sub>B P ?k) \<cdot>\<^sub>B
(P \<mu> \<star>\<^sub>B P \<nu>)"
unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
also have "... = (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, ?g, d ?b \<star>\<^sub>B P ?k] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B ?g, d ?b, P ?k] \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?b) \<star>\<^sub>B P \<nu>)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \<star>\<^sub>B P ?k) \<cdot>\<^sub>B (P \<mu> \<star>\<^sub>B P \<nu>)
= \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \<cdot>\<^sub>B P \<mu> \<star>\<^sub>B P ?k \<cdot>\<^sub>B P \<nu>"
using assms P_def B.interchange by fastforce
also have
"... = ((e ?c \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?b) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?k \<cdot>\<^sub>B P \<nu>"
using assms P_def B.assoc'_naturality [of "e ?c" \<mu> "d ?b"] by simp
also have
"... = ((e ?c \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?b) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P \<nu> \<cdot>\<^sub>B P ?h"
using assms B.comp_arr_dom B.comp_cod_arr B.src_cod B.src_dom
B.trg_cod B.trg_dom P.as_nat_trans.naturality
by simp
also have "... = (((e ?c \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?b) \<star>\<^sub>B P \<nu>) \<cdot>\<^sub>B (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
using assms B.interchange by auto
finally have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \<star>\<^sub>B P ?k) \<cdot>\<^sub>B (P \<mu> \<star>\<^sub>B P \<nu>)
= (((e ?c \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?b) \<star>\<^sub>B P \<nu>) \<cdot>\<^sub>B (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, ?g, d ?b \<star>\<^sub>B P ?k] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?b \<star>\<^sub>B P \<nu>)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B ?f, d ?b, P ?h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
using assms B.assoc_naturality [of "e ?c \<star>\<^sub>B \<mu>" "d ?b" "P \<nu>"] B.comp_assoc by simp
also have "... = (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B ?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?b \<star>\<^sub>B P \<nu>)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, ?f, d ?b \<star>\<^sub>B P ?h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B ?f, d ?b, P ?h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
using assms B.assoc_naturality [of "e ?c" \<mu> "d ?b \<star>\<^sub>B P \<nu>"] B.comp_assoc by simp
also have "... = (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B SRC \<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, ?f, d ?b \<star>\<^sub>B P ?h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B ?f, d ?b, P ?h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
proof -
have
"(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?b \<star>\<^sub>B P \<nu>)
= (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a])"
proof -
have "(e ?c \<star>\<^sub>B ?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?b \<star>\<^sub>B P \<nu>)
= e ?c \<star>\<^sub>B (?g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B d ?b \<star>\<^sub>B P \<nu>)"
using assms P_def B.whisker_left by simp
also have "... = e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a] \<cdot>\<^sub>B (d ?b \<star>\<^sub>B P \<nu>)"
using assms P_def B.comp_cod_arr
B.interchange [of "?g" \<mu> "\<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \<star>\<^sub>B d ?a]" "d ?b \<star>\<^sub>B P \<nu>"]
by simp
also have "... = e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B
(TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a]"
using assms P_def B.assoc'_naturality [of "d ?b" "e ?b" "\<nu> \<star>\<^sub>B d ?a"]
by simp
also have "... = e ?c \<star>\<^sub>B
(\<mu> \<star>\<^sub>B TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (?f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a])"
using assms B.comp_arr_dom
B.interchange [of \<mu> "?f" "TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a""\<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a]"]
by fastforce
also have
"... = (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a])"
using assms B.whisker_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = ((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B LUNIT (?h \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, ?f, d ?b \<star>\<^sub>B P ?h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B ?f, d ?b, P ?h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
proof -
have "((e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a))
= e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B LUNIT (B.cod (\<nu> \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B ((d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a)"
using assms comp_arr_dom B.comp_cod_arr B.whisker_left
B.interchange [of "?g" \<mu> "LUNIT (?k \<star>\<^sub>B d ?a)" "(d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a"]
by fastforce
also have "... = e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B (\<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B LUNIT (?h \<star>\<^sub>B d ?a)"
using assms LUNIT_naturality [of "\<nu> \<star>\<^sub>B d ?a"] by simp
also have "... = (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B ?f \<star>\<^sub>B LUNIT (?h \<star>\<^sub>B d ?a))"
using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left
B.interchange [of \<mu> "?f" "\<nu> \<star>\<^sub>B d ?a" "LUNIT (?h \<star>\<^sub>B d ?a)"]
by simp
finally have "((e ?c \<star>\<^sub>B ?g \<star>\<^sub>B LUNIT (?k \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B TRG \<nu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a))
= (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B ?f \<star>\<^sub>B LUNIT (?h \<star>\<^sub>B d ?a))"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = P (\<mu> \<star>\<^sub>B \<nu>) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B LUNIT (?h \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ?f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, ?f, d ?b \<star>\<^sub>B P ?h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B ?f, d ?b, P ?h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \<star>\<^sub>B P ?h)"
proof -
have "(e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a)
= P (\<mu> \<star>\<^sub>B \<nu>) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a])"
proof -
have "(e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a)
= e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a] \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B \<nu> \<star>\<^sub>B d ?a)"
using assms B.whisker_left by simp
also have "... = e ?c \<star>\<^sub>B ((\<mu> \<star>\<^sub>B \<nu>) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a]"
using assms B.assoc'_naturality [of \<mu> \<nu> "d ?a"] by simp
also have "... = P (\<mu> \<star>\<^sub>B \<nu>) \<cdot>\<^sub>B (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a])"
using assms P_def B.whisker_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P (\<mu> \<star>\<^sub>B \<nu>) \<cdot>\<^sub>B CMP ?f ?h"
unfolding CMP_def LUNIT_def using assms by simp
finally show ?thesis by simp
qed
interpretation EV: self_evaluation_map V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B ..
notation EV.eval ("\<lbrace>_\<rbrace>")
abbreviation (input) SRCt ("\<^bold>S\<^bold>R\<^bold>C")
where "\<^bold>S\<^bold>R\<^bold>C \<mu> \<equiv> \<^bold>\<langle>d (src\<^sub>B \<mu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e (src\<^sub>B \<mu>)\<^bold>\<rangle>"
abbreviation (input) TRGt ("\<^bold>T\<^bold>R\<^bold>G")
where "\<^bold>T\<^bold>R\<^bold>G \<mu> \<equiv> \<^bold>\<langle>d (trg\<^sub>B \<mu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e (trg\<^sub>B \<mu>)\<^bold>\<rangle>"
abbreviation (input) PRJt ("\<^bold>P\<^bold>R\<^bold>J")
where "\<^bold>P\<^bold>R\<^bold>J \<mu> \<equiv> \<^bold>\<langle>e (trg\<^sub>B \<mu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<mu>\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d (src\<^sub>B \<mu>)\<^bold>\<rangle>"
text \<open>
The \<open>CMP f g\<close> satisfy the coherence conditions with respect to associativity that are
required of compositors.
\<close>
lemma CMP_coherence:
assumes "B.ide f" and "B.ide g" and "B.ide h" and "src\<^sub>B f = trg\<^sub>B g" and "src\<^sub>B g = trg\<^sub>B h"
shows "P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B CMP (f \<star>\<^sub>B g) h \<cdot>\<^sub>B (CMP f g \<star>\<^sub>B P h)
= CMP f (g \<star>\<^sub>B h) \<cdot>\<^sub>B (P f \<star>\<^sub>B CMP g h) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]"
proof -
text \<open>
The overall strategy of the proof is to expand the definition of \<open>CMP\<close> on the
left and right-hand sides, then permute the occurrences of \<open>LUNIT\<close> and
\<open>RUNIT\<close> to the left ends of both the left-hand side and right-hand side of the
equation to be proved, so that the initial portions of these expressions become
identical and the remaining parts to the right consist only of canonical isomorphisms.
Then the Coherence Theorem is applied to prove syntactically (and automatically) that the
canonical parts are equal, which implies equality of the complete expressions.
The rest is just grinding through the calculations.
\<close>
let ?a = "trg\<^sub>B f"
let ?b = "trg\<^sub>B g"
let ?c = "trg\<^sub>B h"
let ?d = "src\<^sub>B h"
have "P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B CMP (f \<star>\<^sub>B g) h \<cdot>\<^sub>B (CMP f g \<star>\<^sub>B P h)
= P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g)
\<star>\<^sub>B P h)"
unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
also have "... = P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
using assms B.whisker_right P_def by simp (* 6 sec *)
also have "... = P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B (LUNIT h \<star>\<^sub>B d ?d) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]"
using assms LUNIT_hcomp [of h "d ?d"] B.invert_side_of_triangle by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
using assms B.whisker_left by simp
finally have "e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, h, d ?d]) \<cdot>\<^sub>B (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, h, d ?d] \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B (((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]"
using assms B.assoc'_naturality [of "f \<star>\<^sub>B g" "LUNIT h" "d ?d"] by simp
also have "... = (e ?a \<star>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B (e ?a \<star>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d"
using assms B.whisker_left B.whisker_right P_def by simp
also have "... = e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<cdot>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d"
using assms B.assoc_naturality [of f g "LUNIT h"] by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d)"
using assms B.whisker_left B.whisker_right by simp
finally have "P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B (e ?a \<star>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d)"
by simp
thus ?thesis using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "((e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h)
= (e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h"
using assms TRIANGLE [of f "g \<star>\<^sub>B d ?c"] B.invert_side_of_triangle by simp
also have "... = ((e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h)"
using assms B.whisker_left B.whisker_right P_def by simp
finally have "((e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h)
= ((e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B TRG g, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B ((e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)
= (e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c] \<cdot>\<^sub>B (RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c)) \<star>\<^sub>B P h"
using assms B.whisker_left B.whisker_right P_def by simp
also have "... = (e ?a \<star>\<^sub>B
((RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B TRG g, g, d ?c])
\<star>\<^sub>B P h"
using assms B.assoc'_naturality [of "RUNIT f" g "d ?c"] by simp
also have "... = ((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B TRG g, g, d ?c]) \<star>\<^sub>B P h)"
using assms B.whisker_left B.whisker_right P_def by simp
finally have "((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)
= ((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B TRG g, g, d ?c]) \<star>\<^sub>B P h)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g)) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B ((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)
= \<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<cdot>\<^sub>B (e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h"
using assms B.whisker_left B.whisker_right P_def by simp
also have "... = ((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g)) \<star>\<^sub>B d ?c) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h"
using assms B.assoc'_naturality [of "e ?a" "RUNIT f \<star>\<^sub>B g" "d ?c"] by simp
also have "... = (((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g)) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h)"
using assms B.whisker_left B.whisker_right P_def by simp
finally have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h)
= (((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g)) \<star>\<^sub>B d ?c) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g)) \<star>\<^sub>B d ?c \<star>\<^sub>B P h)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
using assms B.assoc_naturality [of "e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g)" "d ?c" "P h"] B.comp_assoc
by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c \<star>\<^sub>B P h)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
using assms B.assoc_naturality [of "e ?a" "RUNIT f \<star>\<^sub>B g" "d ?c \<star>\<^sub>B P h"] B.comp_assoc
by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d])) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "((e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c \<star>\<^sub>B P h))
= e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d] \<cdot>\<^sub>B (d ?c \<star>\<^sub>B P h)"
using assms B.comp_cod_arr B.whisker_left P_def
B.interchange [of "f \<star>\<^sub>B g" "RUNIT f \<star>\<^sub>B g"]
by simp
also have "... = e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]"
using assms B.comp_arr_dom P_def by simp
finally have "((e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B d ?c \<star>\<^sub>B P h))
= e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) =
e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]"
using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left
B.interchange
[of "f \<star>\<^sub>B g" "RUNIT f \<star>\<^sub>B g" "\<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]" "\<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]"]
by simp (* 3 sec *)
also have "... = (e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d])"
using assms B.comp_arr_dom B.whisker_left
B.interchange [of "RUNIT f \<star>\<^sub>B g" "(f \<star>\<^sub>B SRC f) \<star>\<^sub>B g" "\<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]"
"\<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]"]
by simp
finally have "(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d])
= (e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])
= e ?a \<star>\<^sub>B (\<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B
(f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[f, g, ((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[f, g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B
(f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d])
= \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, g, ((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]"
using assms B.pentagon' B.comp_assoc by simp
moreover have "B.inv (\<a>\<^sub>B\<^sup>-\<^sup>1[f, g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h] \<star>\<^sub>B d ?d)
= \<a>\<^sub>B[f, g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h] \<star>\<^sub>B d ?d"
using assms by simp
ultimately show ?thesis
using assms B.comp_assoc
B.invert_opposite_sides_of_square
[of "\<a>\<^sub>B\<^sup>-\<^sup>1[f, g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h] \<star>\<^sub>B d ?d"
"\<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B
(f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d])"
"\<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d]"
"\<a>\<^sub>B\<^sup>-\<^sup>1[f, g, ((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]"]
by simp (* 3 sec *)
qed
also have "... = (e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, TRG h \<star>\<^sub>B h] \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])
= (e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, (d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])
= e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d] \<cdot>\<^sub>B ((RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B
(RUNIT f \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]"
using assms B.hseqI' B.assoc_naturality [of "RUNIT f" g "\<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]"]
by simp
also have "... = (e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (RUNIT f \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])
= (e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d))) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c \<star>\<^sub>B e ?c, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c \<star>\<^sub>B e ?c, h, d ?d]
= (e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
using assms B.whisker_left B.comp_arr_dom B.comp_cod_arr
B.interchange [of "RUNIT f" "f \<star>\<^sub>B SRC f" "g \<star>\<^sub>B ((TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
"g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]"]
by simp (* 5 sec *)
also have "... = (e ?a \<star>\<^sub>B
(f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
\<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
using assms TRIANGLE [of f "g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d"] by simp
also have "... = (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
using assms B.whisker_left B.comp_assoc by simp
finally have "e ?a \<star>\<^sub>B RUNIT f \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]
= (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B f \<star>\<^sub>B
(LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d ?b \<star>\<^sub>B e ?b, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]"
using assms LUNIT_hcomp [of g "(TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d"]
B.invert_side_of_triangle
by simp
also have "... = (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d])"
using assms B.whisker_left by simp
finally have "e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B f \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B (LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B f \<star>\<^sub>B
((LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]"
using assms B.assoc'_naturality [of "LUNIT g" "TRG h \<star>\<^sub>B h" "d ?d"]
by simp
also have "... = (e ?a \<star>\<^sub>B f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = ((e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B SRC g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B
(f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B
((f \<star>\<^sub>B LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]"
using assms B.assoc'_naturality [of f "LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h" "d ?d"] by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
proof -
have "(e ?a \<star>\<^sub>B (f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d"
using assms B.whisker_left B.whisker_right B.comp_arr_dom B.comp_cod_arr
B.interchange [of g "LUNIT g" "LUNIT h" "(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h"]
by simp (* 5 sec *)
thus ?thesis
using assms by simp
qed
finally have L: "P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B CMP (f \<star>\<^sub>B g) h \<cdot>\<^sub>B (CMP f g \<star>\<^sub>B P h)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
by simp
have "CMP f (g \<star>\<^sub>B h) \<cdot>\<^sub>B (P f \<star>\<^sub>B CMP g h) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]
= (e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT ((g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B
(e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B
(e ?b \<star>\<^sub>B g \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h)) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
also have "... = (e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT ((g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
using assms B.whisker_left P_def B.comp_assoc by auto (* 5 sec *)
also have "... = ((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT ((g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B f \<star>\<^sub>B (LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]"
using assms LUNIT_hcomp [of "g \<star>\<^sub>B h" "d ?d"] B.invert_side_of_triangle
by simp
also have "... = (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT ((g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B (e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B ((f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]"
using assms B.assoc'_naturality [of f "LUNIT (g \<star>\<^sub>B h)" "d ?d"]
LUNIT_in_hom [of "g \<star>\<^sub>B h"]
by auto
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
((P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d)
= P f \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \<star>\<^sub>B d ?d]"
using assms TRIANGLE [of g "h \<star>\<^sub>B d ?d"] B.invert_side_of_triangle by simp
also have "... = (P f \<star>\<^sub>B e ?b \<star>\<^sub>B RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \<star>\<^sub>B d ?d])"
using assms B.whisker_left P_def by simp
finally have "P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B LUNIT (h \<star>\<^sub>B d ?d) =
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \<star>\<^sub>B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B (P f \<star>\<^sub>B e ?b \<star>\<^sub>B RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d)
= P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d] \<cdot>\<^sub>B (RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d)"
using assms B.whisker_left P_def by simp
also have "... = P f \<star>\<^sub>B e ?b \<star>\<^sub>B
((RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]"
using assms B.assoc'_naturality [of "RUNIT g" h "d ?d"] by auto
also have "... = (P f \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d])"
using assms B.whisker_left P_def by simp
finally have "(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B RUNIT g \<star>\<^sub>B h \<star>\<^sub>B d ?d)
= (P f \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
(((e ?a \<star>\<^sub>B f) \<star>\<^sub>B d ?b) \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B (P f \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= \<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d"
using assms B.comp_arr_dom B.comp_cod_arr P_def
B.interchange
[of "\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b]" "P f" "P (g \<star>\<^sub>B h)" "e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d"]
by simp
also have "... = (((e ?a \<star>\<^sub>B f) \<star>\<^sub>B d ?b) \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
using assms B.comp_arr_dom B.comp_cod_arr
B.interchange
[of "(e ?a \<star>\<^sub>B f) \<star>\<^sub>B d ?b" "\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b]"
"e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d"
"e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d"]
by simp (* 4 sec *)
finally have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P (g \<star>\<^sub>B h)) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (((e ?a \<star>\<^sub>B f) \<star>\<^sub>B d ?b) \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P (g \<star>\<^sub>B h)] \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f) \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
using assms P_def B.comp_assoc
B.assoc_naturality [of "e ?a \<star>\<^sub>B f" "d ?b" "e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d"]
by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?c \<star>\<^sub>B e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
using assms P_def B.comp_assoc
B.assoc_naturality [of "e ?a" f "d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d"]
by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have
"(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B f \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d] \<cdot>\<^sub>B
(d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B f \<star>\<^sub>B
(TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]"
using assms B.assoc'_naturality [of "d ?b" "e ?b" "(RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d"] by auto
also have "... = (e ?a \<star>\<^sub>B f \<star>\<^sub>B SRC f \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d])"
using assms B.whisker_left by simp
finally have
"(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B f \<star>\<^sub>B SRC f \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B f \<star>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B f \<star>\<^sub>B
((TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]"
using assms B.assoc'_naturality [of "TRG g" "RUNIT g \<star>\<^sub>B h" "d ?d"] by simp
also have "... = (e ?a \<star>\<^sub>B f \<star>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B f \<star>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d] \<cdot>\<^sub>B
(f \<star>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a \<star>\<^sub>B
((f \<star>\<^sub>B SRC f \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]"
using assms
B.assoc'_naturality [of f "(d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)" "d ?d"]
by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B g \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B (TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d])"
using assms by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B \<a>\<^sub>B[g, d ?c \<star>\<^sub>B e ?c, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d
= e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B(g \<star>\<^sub>B LUNIT h) \<cdot>\<^sub>B \<a>\<^sub>B[g, SRC g, h]) \<star>\<^sub>B d ?d"
using assms TRIANGLE [of g h] by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B \<a>\<^sub>B[g, TRG h, h]) \<star>\<^sub>B d ?d)"
using assms B.whisker_left B.whisker_right by simp
finally have "e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B (RUNIT g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B \<a>\<^sub>B[g, TRG h, h]) \<star>\<^sub>B d ?d)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B \<a>\<^sub>B[g, SRC g, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d
= e ?a \<star>\<^sub>B (f \<star>\<^sub>B (LUNIT g \<star>\<^sub>B h) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \<star>\<^sub>B d ?d"
using assms LUNIT_hcomp [of g h] B.invert_side_of_triangle by simp
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \<star>\<^sub>B d ?d)"
using assms B.whisker_left B.whisker_right by simp
finally have "e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT (g \<star>\<^sub>B h)) \<star>\<^sub>B d ?d
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \<star>\<^sub>B d ?d)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = ((e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B (TRG g \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \<star>\<^sub>B h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B \<a>\<^sub>B[g, SRC g, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have
"(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h] \<cdot>\<^sub>B (TRG g \<star>\<^sub>B g \<star>\<^sub>B LUNIT h)) \<star>\<^sub>B d ?d"
using assms B.whisker_left B.whisker_right by auto
also have "... = e ?a \<star>\<^sub>B
(f \<star>\<^sub>B ((TRG g \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, SRC g \<star>\<^sub>B h]) \<star>\<^sub>B d ?d"
using assms B.assoc'_naturality [of "TRG g" g "LUNIT h"] by auto
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B (TRG g \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \<star>\<^sub>B h]) \<star>\<^sub>B d ?d)"
using assms B.whisker_left B.whisker_right by auto
finally have "(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B (TRG g \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \<star>\<^sub>B h]) \<star>\<^sub>B d ?d)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \<star>\<^sub>B h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B TRG g \<star>\<^sub>B \<a>\<^sub>B[g, d ?c \<star>\<^sub>B e ?c, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
proof -
have "(e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B (TRG g \<star>\<^sub>B g) \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d)
= e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d"
using assms B.whisker_left B.whisker_right B.comp_arr_dom B.comp_cod_arr
B.interchange [of "LUNIT g" "TRG g \<star>\<^sub>B g" h "LUNIT h"]
by simp (* 4 sec *)
thus ?thesis
using assms by simp
qed
finally have R: "CMP f (g \<star>\<^sub>B h) \<cdot>\<^sub>B (P f \<star>\<^sub>B CMP g h) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]
= (e ?a \<star>\<^sub>B (f \<star>\<^sub>B LUNIT g \<star>\<^sub>B LUNIT h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \<star>\<^sub>B h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B \<a>\<^sub>B[g, TRG h, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
using assms by simp
text \<open>
The portions of the expressions on the right-hand sides of assertions \<open>L\<close> and \<open>R\<close>
that are not identical only involve canonical isomorphisms, and thus they can be proved
equal automatically by the simplifier, once we have expressed them in the formal
language of \<open>B\<close>.
\<close>
let ?LHS = "(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \<star>\<^sub>B g) \<star>\<^sub>B TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g \<star>\<^sub>B g, TRG h \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f, d ?b \<star>\<^sub>B e ?b, g \<star>\<^sub>B (TRG h \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B[f \<star>\<^sub>B SRC f, g, TRG h \<star>\<^sub>B h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B ((f \<star>\<^sub>B SRC f) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c \<star>\<^sub>B P h] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c, P h] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \<star>\<^sub>B SRC f) \<star>\<^sub>B g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B SRC f, g, d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
((e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \<star>\<^sub>B d ?c]) \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P g] \<star>\<^sub>B P h) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B P g) \<star>\<^sub>B P h)"
let ?LHSt = "(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>T\<^bold>R\<^bold>G g \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>T\<^bold>R\<^bold>G h \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>T\<^bold>R\<^bold>G h \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g, \<^bold>\<langle>g\<^bold>\<rangle>, (\<^bold>T\<^bold>R\<^bold>G h \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>T\<^bold>R\<^bold>G h \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G h, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>T\<^bold>R\<^bold>G h \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle>, (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J h\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle>, (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
((\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
((\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>S\<^bold>R\<^bold>C f, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?c\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
((\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?c\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J g\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J g\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
((\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J g) \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h)"
let ?RHS = "(e ?a \<star>\<^sub>B (f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \<star>\<^sub>B h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B (f \<star>\<^sub>B SRC f \<star>\<^sub>B \<a>\<^sub>B[g, TRG h, h]) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[f, TRG g \<star>\<^sub>B (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \<star>\<^sub>B SRC g) \<star>\<^sub>B h, d ?d]) \<cdot>\<^sub>B
(e ?a \<star>\<^sub>B f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a, f, d ?b \<star>\<^sub>B P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?a \<star>\<^sub>B f, d ?b, P ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \<star>\<^sub>B e ?b \<star>\<^sub>B ((g \<star>\<^sub>B SRC g) \<star>\<^sub>B h) \<star>\<^sub>B d ?d) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B SRC g, h, d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B e ?b \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \<star>\<^sub>B d ?d]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b, g, d ?c \<star>\<^sub>B P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B g, d ?c, P h]) \<cdot>\<^sub>B
(P f \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \<star>\<^sub>B P h) \<cdot>\<^sub>B
\<a>\<^sub>B[P f, P g, P h]"
let ?RHSt = "(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>T\<^bold>R\<^bold>G h \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C f \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>T\<^bold>R\<^bold>G h, \<^bold>\<langle>h\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>T\<^bold>R\<^bold>G g \<^bold>\<star> (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>\<langle>e ?b\<^bold>\<rangle>, ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>)\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>, (\<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>)\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J f \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>S\<^bold>R\<^bold>C g, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J f \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>S\<^bold>R\<^bold>C g, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J f \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?d\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J f \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J f \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J h\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J f \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?c\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>P\<^bold>R\<^bold>J f, \<^bold>P\<^bold>R\<^bold>J g, \<^bold>P\<^bold>R\<^bold>J h\<^bold>]"
have E: "?LHS = ?RHS"
proof -
have "?LHS = \<lbrace>?LHSt\<rbrace>"
- using assms B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def
+ using assms B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def
by simp
also have "... = \<lbrace>?RHSt\<rbrace>"
using assms by (intro EV.eval_eqI, auto)
also have "... = ?RHS"
- using assms B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def
+ using assms B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def
by simp
finally show ?thesis by blast
qed
show ?thesis
using L R E by argo
qed
interpretation CMP: transformation_by_components B.VV.comp comp HoPP.map PoH.map
\<open>\<lambda>\<mu>\<nu>. CMP (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
proof
show 1: "\<And>fg. B.VV.ide fg \<Longrightarrow> \<guillemotleft>CMP (fst fg) (snd fg) : HoPP.map fg \<Rightarrow> PoH.map fg\<guillemotright>"
- using CMP_in_hom(2) B.VV.ide_char B.VV.arr_char P.FF_def hcomp_def arr_char
+ using CMP_in_hom(2) B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P.FF_def hcomp_def arr_char\<^sub>S\<^sub>b\<^sub>C
P\<^sub>0_props(1) P.preserves_arr
by auto
show "\<And>fg. B.VV.arr fg \<Longrightarrow>
CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) \<cdot> HoPP.map fg =
PoH.map fg \<cdot> CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
proof -
fix fg
assume fg: "B.VV.arr fg"
have "CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) \<cdot> HoPP.map fg =
CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) \<cdot>\<^sub>B HoPP.map fg"
- using fg 1 B.VV.ide_char B.VV.arr_char B.VV.cod_char HoPP.preserves_arr
+ using fg 1 B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C HoPP.preserves_arr
P.FF_def hcomp_char comp_char
by auto
also have "... = PoH.map fg \<cdot>\<^sub>B CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
using fg CMP_naturality [of "fst fg" "snd fg"]
- B.VV.ide_char B.VV.arr_char B.VV.dom_char B.VV.cod_char P.FF_def
- hcomp_def arr_char P\<^sub>0_props(1) P.preserves_arr
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C P.FF_def
+ hcomp_def arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) P.preserves_arr
by auto
also have "... = PoH.map fg \<cdot> CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
proof -
have "P\<^sub>0 (src\<^sub>B (snd fg)) \<in> Obj \<and> P\<^sub>0 (trg\<^sub>B (fst fg)) \<in> Obj"
using fg P\<^sub>0_props(6) B.VV.arrE B.obj_src B.obj_trg by meson
moreover have "B.seq (P (fst fg \<star>\<^sub>B snd fg))
(CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg)))"
- using fg 1 B.VV.ide_char B.VV.arr_char B.VV.dom_char by simp
+ using fg 1 B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C by simp
ultimately show ?thesis
- using fg 1 comp_char arr_char in_hom_char by simp
+ using fg 1 comp_char arr_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
finally show "CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) \<cdot> HoPP.map fg =
PoH.map fg \<cdot> CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
by blast
qed
qed
interpretation CMP: natural_isomorphism B.VV.comp comp HoPP.map PoH.map CMP.map
- using iso_CMP B.VV.ide_char B.VV.arr_char CMP.map_simp_ide
+ using iso_CMP B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C CMP.map_simp_ide
by unfold_locales simp
definition \<Phi>\<^sub>P
where "\<Phi>\<^sub>P = CMP.map"
interpretation \<Phi>\<^sub>P: natural_isomorphism B.VV.comp comp HoPP.map PoH.map \<Phi>\<^sub>P
unfolding \<Phi>\<^sub>P_def
using CMP.natural_isomorphism_axioms by simp
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
lemma \<Phi>\<^sub>P_in_hom [intro]:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "\<guillemotleft>\<Phi>\<^sub>P (f, g) : src (P g) \<rightarrow> trg (P f)\<guillemotright>"
and "\<guillemotleft>\<Phi>\<^sub>P (f, g) : P f \<star> P g \<Rightarrow> P (f \<star>\<^sub>B g)\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<Phi>\<^sub>P (f, g) : P f \<star> P g \<Rightarrow> P (f \<star>\<^sub>B g)\<guillemotright>"
- using assms B.VV.ide_char B.VV.arr_char P.FF_def by auto
+ using assms B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P.FF_def by auto
show "\<guillemotleft>\<Phi>\<^sub>P (f, g) : src (P g) \<rightarrow> trg (P f)\<guillemotright>"
using 1
- by (metis (no_types, lifting) hcomp_simps(2) in_hhom_def in_hom_char
+ by (metis (no_types, lifting) hcomp_simps(2) in_hhom_def in_hom_char\<^sub>S\<^sub>b\<^sub>C
src_hcomp vconn_implies_hpar(1-2))
qed
lemma \<Phi>\<^sub>P_simps [simp]:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g"
shows "arr (\<Phi>\<^sub>P (f, g))"
and "src (\<Phi>\<^sub>P (f, g)) = src (P g)"
and "trg (\<Phi>\<^sub>P (f, g)) = trg (P f)"
and "dom (\<Phi>\<^sub>P (f, g)) = P f \<star> P g"
and "cod (\<Phi>\<^sub>P (f, g)) = P (f \<star>\<^sub>B g)"
using assms \<Phi>\<^sub>P_in_hom by blast+
sublocale prj: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \<a> \<i>\<^sub>B src trg P \<Phi>\<^sub>P
proof
fix f g h
assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h"
and fg: "src\<^sub>B f = trg\<^sub>B g" and gh: "src\<^sub>B g = trg\<^sub>B h"
have 1: "ide (P f) \<and> ide (P g) \<and> ide (P h)"
using f g h P.preserves_ide by simp
have "P \<a>\<^sub>B[f, g, h] \<cdot> \<Phi>\<^sub>P (f \<star>\<^sub>B g, h) \<cdot> (\<Phi>\<^sub>P (f, g) \<star> P h) =
P \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B \<Phi>\<^sub>P (f \<star>\<^sub>B g, h) \<cdot>\<^sub>B (\<Phi>\<^sub>P (f, g) \<star>\<^sub>B P h)"
- using f g h fg gh B.VV.arr_char B.VV.dom_char B.VV.cod_char P.FF_def
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C P.FF_def
by (intro comp_eqI hcomp_eqI seqI hseqI') auto
also have "... = CMP f (g \<star>\<^sub>B h) \<cdot>\<^sub>B (P f \<star>\<^sub>B CMP g h) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]"
- using f g h fg gh CMP_coherence CMP.map_simp_ide B.VV.ide_char B.VV.arr_char
+ using f g h fg gh CMP_coherence CMP.map_simp_ide B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
\<Phi>\<^sub>P_def
by simp
also have "... = \<Phi>\<^sub>P (f, g \<star>\<^sub>B h) \<cdot>\<^sub>B (P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]"
- using f g h fg gh B.VV.ide_char B.VV.arr_char \<Phi>\<^sub>P_def by simp
+ using f g h fg gh B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>P_def by simp
also have "... = \<Phi>\<^sub>P (f, g \<star>\<^sub>B h) \<cdot> (P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) \<cdot> \<a>\<^sub>B[P f, P g, P h]"
proof -
have 2: "arr ((P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h])"
- using f g h fg gh B.VV.arr_char B.VV.dom_char P.FF_def by auto
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C P.FF_def by auto
moreover have "(P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h] =
(P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]"
using f g h fg gh 2 comp_eqI hcomp_eqI assoc_simp by auto
moreover have 3: "arr (P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h))"
using f g h fg gh arr_hcompI
by (intro arr_hcompI hseqI') auto
moreover have "B.dom (P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) = P f \<star>\<^sub>B P g \<star>\<^sub>B P h"
proof -
have "B.dom (P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) = P f \<star>\<^sub>B B.dom (\<Phi>\<^sub>P (g, h))"
using f g h fg gh 3 B.hcomp_simps(3)
- by (metis (no_types, lifting) 1 arrE ideD(1) ideD(2) dom_char)
+ by (metis (no_types, lifting) 1 arrE ideD(1) ideD(2) dom_char\<^sub>S\<^sub>b\<^sub>C)
also have "... = P f \<star>\<^sub>B P g \<star>\<^sub>B P h"
- using g h gh dom_char hcomp_char \<Phi>\<^sub>P_simps(1-4) by auto
+ using g h gh dom_char\<^sub>S\<^sub>b\<^sub>C hcomp_char \<Phi>\<^sub>P_simps(1-4) by auto
finally show ?thesis by blast
qed
moreover have "B.dom (\<Phi>\<^sub>P (f, g \<star>\<^sub>B h)) =
B.cod ((P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h])"
proof -
have "B.dom (\<Phi>\<^sub>P (f, g \<star>\<^sub>B h)) = dom (\<Phi>\<^sub>P (f, g \<star>\<^sub>B h))"
- using f g h fg gh B.VV.arr_char dom_char by simp
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = cod ((P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h])"
- using f g h fg gh 2 VV.arr_char \<Phi>\<^sub>P_simps by simp
+ using f g h fg gh 2 VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>P_simps by simp
also have "... = B.cod ((P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h])"
- using 2 cod_char by presburger
+ using 2 cod_char\<^sub>S\<^sub>b\<^sub>C by presburger
also have "... = B.cod ((P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h])"
proof -
have "(P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h] = (P f \<star>\<^sub>B \<Phi>\<^sub>P (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[P f, P g, P h]"
using f g h fg gh 2 comp_eqI hcomp_eqI assoc_simp by auto
thus ?thesis by presburger
qed
finally show ?thesis by blast
qed
moreover have "B.cod \<a>\<^sub>B[P f, P g, P h] = P f \<star>\<^sub>B P g \<star>\<^sub>B P h"
- using f g h fg gh 1 ide_char by auto
+ using f g h fg gh 1 ide_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show ?thesis
- using f g h fg gh B.VV.arr_char assoc_simp assoc_simps(1) dom_char cod_char
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C assoc_simp assoc_simps(1) dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C
by (intro comp_eqI' seqI arr_compI arr_hcompI) auto
qed
also have "... = \<Phi>\<^sub>P (f, g \<star>\<^sub>B h) \<cdot> (P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h]"
using f g h fg gh assoc_simp hcomp_eqI \<Phi>\<^sub>P_simps(1) by auto
finally show "P \<a>\<^sub>B[f, g, h] \<cdot> \<Phi>\<^sub>P (f \<star>\<^sub>B g, h) \<cdot> (\<Phi>\<^sub>P (f, g) \<star> P h) =
\<Phi>\<^sub>P (f, g \<star>\<^sub>B h) \<cdot> (P f \<star> \<Phi>\<^sub>P (g, h)) \<cdot> \<a>[P f, P g, P h]"
by blast
qed
lemma pseudofunctor_prj:
shows "pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \<a> \<i>\<^sub>B src trg P \<Phi>\<^sub>P"
..
text \<open>
We need an explicit formula for the unit constraints for \<open>P\<close>.
\<close>
lemma prj_unit_char:
assumes "B.obj a"
shows "prj.unit a = (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)"
proof -
text \<open>
We eventually need one of the triangle identities from the following interpretation.
However in the meantime, it contains a lot of simps that make an otherwise arduous
calculation much easier. So interpret it up front.
\<close>
interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B \<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence(1) by simp
let ?x = "(e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)"
have x: "\<guillemotleft>?x : P.map\<^sub>0 a \<Rightarrow>\<^sub>B P a\<guillemotright>"
using assms P_def P.map\<^sub>0_def P_map\<^sub>0_simp
by (intro B.comp_in_homI') auto
have "?x = prj.unit a"
proof (intro prj.unit_eqI)
show "B.obj a" by fact
have a_da: "\<guillemotleft>a \<star>\<^sub>B d a : P\<^sub>0 a \<rightarrow>\<^sub>B a\<guillemotright> \<and> B.ide (a \<star>\<^sub>B d a)"
using assms B.obj_simps by auto
show x\<^sub>S: "\<guillemotleft>?x : P.map\<^sub>0 a \<Rightarrow> P a\<guillemotright>"
- using assms x P_map\<^sub>0_simp P_simps\<^sub>B(1) arr_char B.arrI
+ using assms x P_map\<^sub>0_simp P_simps\<^sub>B(1) arr_char\<^sub>S\<^sub>b\<^sub>C B.arrI
equivalence_data_simps\<^sub>B(6) counit_simps(1,4) B.obj_simps(1)
B.src.preserves_arr B.vconn_implies_hpar(1-4)
- by (metis (no_types, lifting) P_simps(1) in_hom_char)
+ by (metis (no_types, lifting) P_simps(1) in_hom_char\<^sub>S\<^sub>b\<^sub>C)
show "iso ?x"
- using assms x\<^sub>S B.isos_compose iso_char arr_char by auto
+ using assms x\<^sub>S B.isos_compose iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by auto
have *: "\<guillemotleft>?x : P\<^sub>0 a \<rightarrow> P\<^sub>0 a\<guillemotright>"
using assms x\<^sub>S P\<^sub>0_props vconn_implies_hpar(1-2)
by (intro in_hhomI) auto
show "?x \<cdot> \<i>\<^sub>B[P.map\<^sub>0 a] = (P \<i>\<^sub>B[a] \<cdot> \<Phi>\<^sub>P (a, a)) \<cdot> (?x \<star> ?x)"
proof -
have 0: "\<guillemotleft>d a \<star>\<^sub>B e a \<star>\<^sub>B a \<star>\<^sub>B d a : P\<^sub>0 a \<rightarrow>\<^sub>B a\<guillemotright>"
using assms by force
have 1: "B.seq (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"
using assms by (elim B.objE, intro B.seqI) auto
have "(P \<i>\<^sub>B[a] \<cdot> \<Phi>\<^sub>P (a, a)) \<cdot> (?x \<star> ?x) = P \<i>\<^sub>B[a] \<cdot> \<Phi>\<^sub>P (a, a) \<cdot> (?x \<star> ?x)"
using comp_assoc by simp
also have "... = P \<i>\<^sub>B[a] \<cdot> (CMP a a \<cdot> (P a \<star> P a)) \<cdot> (?x \<star> ?x)"
unfolding \<Phi>\<^sub>P_def CMP.map_def
- using assms B.VV.arr_char B.VV.cod_char P.FF_def by auto
+ using assms B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C P.FF_def by auto
also have "... = P \<i>\<^sub>B[a] \<cdot> (CMP a a \<cdot> (P a \<star>\<^sub>B P a)) \<cdot> (?x \<star>\<^sub>B ?x)"
using assms x\<^sub>S hcomp_char src_def trg_def by auto
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B (CMP a a \<cdot>\<^sub>B (P a \<star>\<^sub>B P a)) \<cdot>\<^sub>B (?x \<star>\<^sub>B ?x)"
proof -
have "\<guillemotleft>P \<i>\<^sub>B[a] \<cdot> (CMP a a \<cdot> (P a \<star>\<^sub>B P a)) \<cdot> (?x \<star>\<^sub>B ?x) : P\<^sub>0 a \<star> P\<^sub>0 a \<Rightarrow> P a\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>?x \<star>\<^sub>B ?x : P\<^sub>0 a \<star> P\<^sub>0 a \<Rightarrow> P a \<star> P a\<guillemotright>"
proof -
have "\<guillemotleft>?x \<star> ?x : P\<^sub>0 a \<star> P\<^sub>0 a \<Rightarrow> P a \<star> P a\<guillemotright>"
using assms x\<^sub>S * P\<^sub>0_props P.map\<^sub>0_simps(3)
by (intro hcomp_in_vhom) auto
moreover have "?x \<star> ?x = ?x \<star>\<^sub>B ?x"
using x\<^sub>S * by (intro hcomp_eqI) auto
ultimately show ?thesis by simp
qed
show "\<guillemotleft>P a \<star>\<^sub>B P a : P a \<star> P a \<Rightarrow> P a \<star> P a\<guillemotright>"
using assms hcomp_eqI by fastforce
show "\<guillemotleft>CMP a a : P a \<star> P a \<Rightarrow> P (a \<star>\<^sub>B a)\<guillemotright>"
using assms CMP_in_hom(2) by auto
show "\<guillemotleft>P \<i>\<^sub>B[a] : P (a \<star>\<^sub>B a) \<Rightarrow> P a\<guillemotright>"
using assms by auto
qed
thus ?thesis
by (intro comp_eqI hcomp_eqI) auto
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a] \<cdot>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a)) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
(P a \<star>\<^sub>B P a) \<cdot>\<^sub>B
(?x \<star>\<^sub>B ?x)"
using assms B.comp_assoc CMP_def by auto
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
(P a \<star>\<^sub>B P a) \<cdot>\<^sub>B
(?x \<star>\<^sub>B ?x)"
using assms 1 B.whisker_left B.comp_assoc by fastforce
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((P a \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]))) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
using assms B.interchange B.comp_assoc by simp
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
((e a \<star>\<^sub>B a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "(P a \<star>\<^sub>B P a) \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]))
= ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]))"
unfolding P_def
using assms B.comp_cod_arr [of "(e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
"(e a \<star>\<^sub>B a \<star>\<^sub>B d a) \<star>\<^sub>B (e a \<star>\<^sub>B a \<star>\<^sub>B d a)"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
((e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "(e a \<star>\<^sub>B a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
= e a \<star>\<^sub>B (a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B (a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
proof -
have "B.seq (a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) (a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
using assms by (elim B.objE, intro B.seqI) auto
thus ?thesis
using assms B.whisker_left by simp
qed
also have "... = e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"
using assms 1 B.whisker_left by fastforce
finally have "(e a \<star>\<^sub>B a \<star>\<^sub>B B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
= e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a)) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "(e a \<star>\<^sub>B a \<star>\<^sub>B
(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a]
= \<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
using assms 1 a_da P_def
B.assoc_naturality
[of "e a" a "(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"]
by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a \<star>\<^sub>B P a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have 1: "B.ide (e a) \<and> B.ide a \<and> B.ide (d a) \<and> B.ide (P a)"
- using assms ide_char P.preserves_ide by auto
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide by auto
have 2: "src\<^sub>B (e a) = trg\<^sub>B a \<and> src\<^sub>B a = trg\<^sub>B (d a) \<and> src\<^sub>B (d a) = trg\<^sub>B (P a)"
using assms by auto
have "((e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]) \<cdot>\<^sub>B (\<a>\<^sub>B[e a, a, d a] \<star>\<^sub>B P a)
= \<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B \<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a]"
using assms 1 2 B.pentagon B.comp_assoc by force
hence "(e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]
= (\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B \<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a]) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a)"
using assms 1 2
B.invert_side_of_triangle(2)
[of "\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B \<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a]"
"(e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]"
"\<a>\<^sub>B[e a, a, d a] \<star>\<^sub>B P a"]
by fastforce
hence "\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B (\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a)
= \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]"
using assms 1 2 B.invert_side_of_triangle(1) B.comp_assoc by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "B.arr (e a) \<and> B.arr a"
using assms by auto
moreover have "B.dom (e a) = e a \<and> B.dom a = a \<and> B.cod a = a \<and>
B.cod (e a) = e a"
using assms by auto
moreover have "B.seq (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"
using assms by (elim B.objE, intro B.seqI) auto
moreover have "src\<^sub>B a =
trg\<^sub>B ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
using assms a_da trg_vcomp by fastforce
moreover have "src\<^sub>B a = a \<and> trg\<^sub>B a = a"
using assms by auto
moreover have "B.dom ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
= d a \<star>\<^sub>B e a \<star>\<^sub>B a \<star>\<^sub>B d a"
using assms a_da by fastforce
moreover have "B.cod ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
= a \<star>\<^sub>B a \<star>\<^sub>B d a"
using assms a_da by fastforce
ultimately have "((e a \<star>\<^sub>B a) \<star>\<^sub>B
(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a \<star>\<^sub>B P a]
= \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
unfolding P_def
using assms
B.assoc'_naturality
[of "e a" a "(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have 1: "B.ide (e a) \<and> B.ide a \<and> B.ide (d a) \<and> B.ide (P a)"
- using assms ide_char P.preserves_ide by auto
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide by auto
have 2: "src\<^sub>B (e a) = trg\<^sub>B a \<and> src\<^sub>B a = trg\<^sub>B (d a) \<and> src\<^sub>B (d a) = trg\<^sub>B (P a)"
using assms by auto
have "((e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]) \<cdot>\<^sub>B (\<a>\<^sub>B[e a, a, d a] \<star>\<^sub>B P a)
= \<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B \<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a]"
using assms 1 2 B.pentagon B.comp_assoc by fastforce
hence "(e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]
= \<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B \<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a)"
using assms 1 2 P.preserves_ide B.comp_assoc
B.invert_side_of_triangle(2)
[of "\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a] \<cdot>\<^sub>B \<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a]"
"(e a \<star>\<^sub>B \<a>\<^sub>B[a, d a, P a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a \<star>\<^sub>B d a, P a]"
"\<a>\<^sub>B[e a, a, d a] \<star>\<^sub>B P a"]
by force
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "(e a \<star>\<^sub>B a \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, d a \<star>\<^sub>B P a]
= \<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B ((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
proof -
have 1: "B.seq (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"
using assms by (elim B.objE, intro B.seqI) auto
moreover have "src\<^sub>B a =
trg\<^sub>B ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])"
using a_da by force
moreover have "B.arr a \<and> B.dom a = a \<and> B.cod a = a \<and> src\<^sub>B a = a \<and> trg\<^sub>B a = a"
using assms by auto
moreover have "B.dom ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
= d a \<star>\<^sub>B e a \<star>\<^sub>B a \<star>\<^sub>B d a"
using a_da by auto
moreover have "B.cod ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
= a \<star>\<^sub>B a \<star>\<^sub>B d a"
using a_da by auto
ultimately show ?thesis
unfolding P_def
using assms
B.assoc_naturality
[of "e a" a "(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]"]
by auto
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B
B.inv (\<epsilon> a))) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))
= ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B B.inv (\<epsilon> a))
= (((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
((e a \<star>\<^sub>B d a) \<star>\<^sub>B B.inv (\<epsilon> a))) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
using assms B.comp_arr_dom B.comp_cod_arr B.comp_assoc
B.interchange [of "e a \<star>\<^sub>B d a" "B.inv (\<epsilon> a)" "B.inv (\<epsilon> a)" "P\<^sub>0 a"]
by simp
also have "... = ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B (e a \<star>\<^sub>B d a) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B
B.inv (\<epsilon> a)) \<cdot>\<^sub>B (B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
using assms B.interchange by simp
also have "... = ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
using assms B.comp_arr_dom by simp
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B
B.inv (\<epsilon> a))
= (\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P a \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B
B.inv (\<epsilon> a))"
proof -
have "B.seq \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using assms a_da by (elim B.objE, intro B.seqI) auto
moreover have "B.seq (P a) ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))"
using assms a_da P_def by auto
ultimately show ?thesis
using assms B.interchange by simp
qed
also have "... = (\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B
B.inv (\<epsilon> a))"
proof -
have "P a \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) = (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)"
using assms x P_def B.comp_cod_arr by blast
moreover have "\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) = \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a"
proof -
have "\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])
= B.inv ((e a \<star>\<^sub>B \<l>\<^sub>B[d a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, d a])"
proof -
have "B.inv ((e a \<star>\<^sub>B \<l>\<^sub>B[d a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, d a])
= B.inv \<a>\<^sub>B[e a, a, d a] \<cdot>\<^sub>B B.inv (e a \<star>\<^sub>B \<l>\<^sub>B[d a])"
proof -
have "B.seq (e a \<star>\<^sub>B \<l>\<^sub>B[d a]) \<a>\<^sub>B[e a, a, d a]"
proof -
have "B.iso ((e a \<star>\<^sub>B \<l>\<^sub>B[d a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, d a])"
using assms by (elim B.objE, intro B.isos_compose) auto
thus ?thesis
by blast
qed
moreover have "B.iso \<a>\<^sub>B[e a, a, d a]"
using assms by force
ultimately show ?thesis
using assms B.inv_comp by auto
qed
also have "... = \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using assms
by (elim B.objE) (simp add: assms)
finally show ?thesis by simp
qed
also have "... = B.inv (\<r>\<^sub>B[e a] \<star>\<^sub>B d a)"
using assms B.triangle [of "d a" "e a"] by simp
also have "... = \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a"
using assms by simp
finally show ?thesis by blast
qed
ultimately show ?thesis by simp
qed
also have "... = (((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]) \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B
(e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<cdot>\<^sub>B P\<^sub>0 a"
using assms 0 B.comp_cod_arr B.comp_arr_dom
by (elim B.objE) auto
also have "... = ((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B P\<^sub>0 a"
using B.comp_assoc by simp
also have "... = (((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "B.seq ((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) (\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]))"
using assms 0 a_da
by (elim B.objE, intro B.seqI) auto
moreover have "B.seq ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) (P\<^sub>0 a)"
using assms 0
apply (intro B.seqI, auto simp add: B.obj_simps(5) P\<^sub>0_props(5))
by (meson B.in_hhomE B.obj_simps(1) a_da)
ultimately show ?thesis
using assms B.interchange by blast
qed
finally have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B P a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))
= (((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a)"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P a] \<cdot>\<^sub>B
(((e a \<star>\<^sub>B a) \<star>\<^sub>B d a) \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))
= ((e a \<star>\<^sub>B a) \<star>\<^sub>B d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a]"
proof -
have "B.hseq (e a) a"
using assms by force
moreover have "src\<^sub>B (d a) = trg\<^sub>B ?x"
using assms B.trg_vcomp [of "e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]" "B.inv (\<epsilon> a)"] by simp
moreover have "B.cod ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) = P a"
using assms x by blast
ultimately show ?thesis
using assms B.assoc_naturality [of "e a \<star>\<^sub>B a" "d a" ?x] by auto
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B
(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "((e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))
= (e a \<star>\<^sub>B a) \<star>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "B.ide (e a \<star>\<^sub>B a)"
using assms by force
moreover have "B.seq ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a])
(d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))"
using a_da B.whisker_left by auto
ultimately show ?thesis
using assms B.whisker_left B.comp_assoc by fastforce
qed
thus ?thesis by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B
(B.inv (\<eta> a) \<star>\<^sub>B d a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B B.inv (\<epsilon> a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))
= (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "(B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))
= (B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B (\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B (d a \<star>\<^sub>B B.inv (\<epsilon> a))"
using assms B.whisker_left B.comp_assoc by simp
also have "... = ((B.inv (\<eta> a) \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B ((d a \<star>\<^sub>B e a) \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B (d a \<star>\<^sub>B B.inv (\<epsilon> a))"
using assms B.assoc'_naturality [of "d a" "e a" "\<l>\<^sub>B\<^sup>-\<^sup>1[d a]"] B.comp_assoc by simp
also have "... = (B.inv (\<eta> a) \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B B.inv (\<epsilon> a))"
using assms B.interchange [of "B.inv (\<eta> a)" "d a \<star>\<^sub>B e a" "a \<star>\<^sub>B d a" "\<l>\<^sub>B\<^sup>-\<^sup>1[d a]"]
B.comp_arr_dom B.comp_cod_arr
by simp
also have "... = (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B (B.inv (\<eta> a) \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B
(d a \<star>\<^sub>B B.inv (\<epsilon> a))"
using assms B.interchange [of a "B.inv (\<eta> a)" "\<l>\<^sub>B\<^sup>-\<^sup>1[d a]" "d a"]
B.comp_arr_dom B.comp_cod_arr B.comp_assoc
by simp
finally show ?thesis by blast
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<r>\<^sub>B[d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a)"
proof -
have "(B.inv (\<eta> a) \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B (d a \<star>\<^sub>B B.inv (\<epsilon> a))
= \<l>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<r>\<^sub>B[d a]"
proof -
have "(B.inv (\<eta> a) \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \<cdot>\<^sub>B (d a \<star>\<^sub>B B.inv (\<epsilon> a))
= B.inv (\<eta> a \<star>\<^sub>B d a) \<cdot>\<^sub>B B.inv \<a>\<^sub>B[d a, e a, d a] \<cdot>\<^sub>B B.inv (d a \<star>\<^sub>B \<epsilon> a)"
using assms by simp
also have "... = B.inv ((d a \<star>\<^sub>B \<epsilon> a) \<cdot>\<^sub>B \<a>\<^sub>B[d a, e a, d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B d a))"
proof -
have "B.iso ((d a \<star>\<^sub>B \<epsilon> a) \<cdot>\<^sub>B \<a>\<^sub>B[d a, e a, d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B d a))"
using assms by (intro B.isos_compose) auto
thus ?thesis
using assms B.inv_comp_left B.comp_assoc by auto
qed
also have "... = B.inv (\<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a])"
using assms triangle_right by simp
also have "... = \<l>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<r>\<^sub>B[d a]"
using assms B.inv_comp by simp
finally show ?thesis by blast
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<r>\<^sub>B[d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
\<r>\<^sub>B\<^sup>-\<^sup>1[e a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
B.inv (\<epsilon> a) \<cdot>\<^sub>B
\<i>\<^sub>B[P\<^sub>0 a]"
proof -
have "\<r>\<^sub>B[e a \<star>\<^sub>B d a] \<cdot>\<^sub>B (B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a) = B.inv (\<epsilon> a) \<cdot>\<^sub>B \<i>\<^sub>B[P\<^sub>0 a]"
using assms B.runit_naturality [of "B.inv (\<epsilon> a)"] B.unitor_coincidence P\<^sub>0_props
by simp
hence "B.inv (\<epsilon> a) \<star>\<^sub>B P\<^sub>0 a = \<r>\<^sub>B\<^sup>-\<^sup>1[e a \<star>\<^sub>B d a] \<cdot>\<^sub>B B.inv (\<epsilon> a) \<cdot>\<^sub>B \<i>\<^sub>B[P\<^sub>0 a]"
using assms P\<^sub>0_props(5) B.invert_side_of_triangle(1) by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<cdot>\<^sub>B \<i>\<^sub>B[P\<^sub>0 a]"
proof -
have P\<^sub>0_a: "B.obj (P\<^sub>0 a) \<and> B.arr (P\<^sub>0 a)"
using assms a_da by fastforce
have i_a: "B.\<rr> a = \<i>\<^sub>B[a]"
using assms B.unitor_coincidence B.\<rr>_ide_simp B.obj_simps
by (metis (no_types, lifting) B.objE)
have "P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<r>\<^sub>B[d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
\<r>\<^sub>B\<^sup>-\<^sup>1[e a \<star>\<^sub>B d a]
= e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]"
proof -
have "P \<i>\<^sub>B[a] \<cdot>\<^sub>B
(e a \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \<cdot>\<^sub>B
(e a \<star>\<^sub>B a \<star>\<^sub>B \<l>\<^sub>B[a \<star>\<^sub>B d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a \<star>\<^sub>B a \<star>\<^sub>B d a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<r>\<^sub>B[d a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, P\<^sub>0 a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
\<r>\<^sub>B\<^sup>-\<^sup>1[e a \<star>\<^sub>B d a]
= \<lbrace>(\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<r>\<^bold>[\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0\<^bold>] \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<l>\<^bold>[\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e a\<^bold>\<rangle>, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e a\<^bold>\<rangle>, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e a\<^bold>\<rangle>, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0) \<^bold>\<star> (\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<r>\<^bold>[\<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>d a\<^bold>\<rangle>, \<^bold>\<langle>P\<^sub>0 a\<^bold>\<rangle>\<^sub>0\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e a\<^bold>\<rangle>, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>P\<^sub>0 a\<^bold>\<rangle>\<^sub>0) \<^bold>\<cdot>
\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]\<rbrace>"
unfolding P_def
using assms B.obj_def [of a] P\<^sub>0_a i_a B.\<alpha>_def B.\<alpha>'.map_ide_simp
- B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char
+ B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
B.\<ll>_ide_simp B.\<rr>_ide_simp
by (elim B.objE) simp
also have "... = \<lbrace>\<^bold>\<langle>e a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]\<rbrace>"
using assms P\<^sub>0_a by (intro EV.eval_eqI) auto
also have "... = e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]"
using assms P_def B.\<ll>_ide_simp by simp
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = ?x \<cdot> \<i>\<^sub>B[P.map\<^sub>0 a]"
proof -
have "B.arr \<i>\<^sub>B[P.map\<^sub>0 a] \<and> src\<^sub>B \<i>\<^sub>B[P.map\<^sub>0 a] \<in> Obj \<and> trg\<^sub>B \<i>\<^sub>B[P.map\<^sub>0 a] \<in> Obj \<and>
P\<^sub>0 a \<cdot>\<^sub>B P\<^sub>0 a \<in> Obj \<and> B.cod \<i>\<^sub>B[P.map\<^sub>0 a] = P\<^sub>0 a"
- using assms P\<^sub>0_props arr_char unit_simps(1)
+ using assms P\<^sub>0_props arr_char\<^sub>S\<^sub>b\<^sub>C unit_simps(1)
apply auto
using obj_char
by (metis (no_types, lifting) B.comp_ide_self B.objE)
thus ?thesis
using assms comp_def B.comp_assoc P\<^sub>0_props [of a] by simp
qed
finally show ?thesis by argo
qed
qed
thus ?thesis by simp
qed
sublocale PoE: composite_pseudofunctor
comp hcomp \<a> \<i>\<^sub>B src trg V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
comp hcomp \<a> \<i>\<^sub>B src trg
E \<Phi>\<^sub>E P \<Phi>\<^sub>P
..
sublocale EoP: composite_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \<a> \<i>\<^sub>B src trg V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
P \<Phi>\<^sub>P E \<Phi>\<^sub>E
..
sublocale I\<^sub>C: identity_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B ..
sublocale I\<^sub>S: identity_pseudofunctor comp hcomp \<a> \<i>\<^sub>B src trg ..
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
abbreviation (input) unit\<^sub>0
where "unit\<^sub>0 \<equiv> e"
definition unit\<^sub>1
where "unit\<^sub>1 f = (e (trg\<^sub>B f) \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e (trg\<^sub>B f), f, src\<^sub>B f] \<cdot>\<^sub>B
((e (trg\<^sub>B f) \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> (src\<^sub>B f))) \<cdot>\<^sub>B
\<a>\<^sub>B[e (trg\<^sub>B f) \<star>\<^sub>B f, d (src\<^sub>B f), e (src\<^sub>B f)] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e (trg\<^sub>B f), f, d (src\<^sub>B f)] \<star>\<^sub>B e (src\<^sub>B f))"
abbreviation (input) \<eta>\<^sub>0
where "\<eta>\<^sub>0 \<equiv> unit\<^sub>0"
abbreviation (input) \<eta>\<^sub>1
where "\<eta>\<^sub>1 \<equiv> unit\<^sub>1"
lemma unit\<^sub>1_in_hom [intro]:
assumes "B.ide f"
shows "\<guillemotleft>\<eta>\<^sub>1 f : src\<^sub>B f \<rightarrow>\<^sub>B P\<^sub>0 (trg\<^sub>B f)\<guillemotright>"
and "\<guillemotleft>\<eta>\<^sub>1 f : (e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f)) \<star>\<^sub>B e (src\<^sub>B f) \<Rightarrow>\<^sub>B e (trg\<^sub>B f) \<star>\<^sub>B f\<guillemotright>"
proof -
show "\<guillemotleft>\<eta>\<^sub>1 f : (e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f)) \<star>\<^sub>B e (src\<^sub>B f) \<Rightarrow>\<^sub>B e (trg\<^sub>B f) \<star>\<^sub>B f\<guillemotright>"
using assms unit\<^sub>1_def by force
thus "\<guillemotleft>\<eta>\<^sub>1 f : src\<^sub>B f \<rightarrow>\<^sub>B P\<^sub>0 (trg\<^sub>B f)\<guillemotright>"
using assms B.vconn_implies_hpar(1-2) by auto
qed
lemma unit\<^sub>1_simps [simp]:
assumes "B.ide f"
shows "B.arr (\<eta>\<^sub>1 f)"
and "src\<^sub>B (\<eta>\<^sub>1 f) = src\<^sub>B f" and "trg\<^sub>B (\<eta>\<^sub>1 f) = P\<^sub>0 (trg\<^sub>B f)"
and "B.dom (\<eta>\<^sub>1 f) = (e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f)) \<star>\<^sub>B e (src\<^sub>B f)"
and "B.cod (\<eta>\<^sub>1 f) = e (trg\<^sub>B f) \<star>\<^sub>B f"
and "B.iso (\<eta>\<^sub>1 f)"
using assms unit\<^sub>1_in_hom
apply auto
unfolding unit\<^sub>1_def
by (intro B.isos_compose) auto (* 10 sec *)
lemma unit\<^sub>1_in_hom\<^sub>S [intro]:
assumes "ide f"
shows "\<guillemotleft>\<eta>\<^sub>1 f : src f \<rightarrow> P\<^sub>0 (trg f)\<guillemotright>"
and "\<guillemotleft>\<eta>\<^sub>1 f : PoE.map f \<star> e (src f) \<Rightarrow> e (trg f) \<star> I\<^sub>S.map f\<guillemotright>"
- using assms ide_char arr_char P\<^sub>0_props(1,6) P.preserves_ide hcomp_def
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1,6) P.preserves_ide hcomp_def
src_def trg_def P_def emb.map_def equivalence_data_simps(2)
- in_hom_char
+ in_hom_char\<^sub>S\<^sub>b\<^sub>C
by auto
lemma unit\<^sub>1_simps\<^sub>S [simp]:
assumes "ide f"
shows "arr (\<eta>\<^sub>1 f)"
and "src (\<eta>\<^sub>1 f) = src f" and "trg (\<eta>\<^sub>1 f) = P\<^sub>0 (trg f)"
and "dom (\<eta>\<^sub>1 f) = PoE.map f \<star> e (src f)"
and "cod (\<eta>\<^sub>1 f) = e (trg f) \<star> I\<^sub>S.map f"
and "iso (\<eta>\<^sub>1 f)"
- using assms unit\<^sub>1_in_hom\<^sub>S iso_char ide_char arr_char P.as_nat_iso.components_are_iso
+ using assms unit\<^sub>1_in_hom\<^sub>S iso_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C P.as_nat_iso.components_are_iso
by auto
sublocale unit: pseudonatural_equivalence comp hcomp \<a> \<i>\<^sub>B src trg
comp hcomp \<a> \<i>\<^sub>B src trg
I\<^sub>S.map I\<^sub>S.cmp \<open>P \<circ> E\<close> PoE.cmp unit\<^sub>0 unit\<^sub>1
proof
show "\<And>a. obj a \<Longrightarrow> ide (e a)"
- using obj_char ide_char arr_char P\<^sub>0_props(1) by force
+ using obj_char ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) by force
show "\<And>a. obj a \<Longrightarrow> \<guillemotleft>e a : src (I\<^sub>S.map a) \<rightarrow> src ((P \<circ> E) a)\<guillemotright>"
using emb.map_def
apply (intro in_hhomI)
apply auto
using obj_char by auto
show "\<And>a. obj a \<Longrightarrow> equivalence_map (e a)"
proof -
fix a
assume a: "obj a"
interpret Adj': equivalence_in_bicategory comp hcomp \<a> \<i>\<^sub>B src trg
\<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using a by unfold_locales auto
show "equivalence_map (e a)"
using Adj'.equivalence_in_bicategory_axioms equivalence_map_def by auto
qed
show "\<And>f. ide f \<Longrightarrow> \<guillemotleft>\<eta>\<^sub>1 f : PoE.map f \<star> e (src f) \<Rightarrow> e (trg f) \<star> I\<^sub>S.map f\<guillemotright>"
using unit\<^sub>1_in_hom\<^sub>S(2) by simp
show "\<And>f. ide f \<Longrightarrow> iso (\<eta>\<^sub>1 f)"
by simp
show *: "\<And>\<mu>. arr \<mu> \<Longrightarrow>
\<eta>\<^sub>1 (cod \<mu>) \<cdot> (PoE.map \<mu> \<star> e (src \<mu>)) = (e (trg \<mu>) \<star> I\<^sub>S.map \<mu>) \<cdot> \<eta>\<^sub>1 (dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
have 1: "B.arr \<mu>"
- using \<mu> arr_char by simp
+ using \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C by simp
let ?f = "B.dom \<mu>"
let ?g = "B.cod \<mu>"
let ?a = "src\<^sub>B \<mu>"
let ?b = "trg\<^sub>B \<mu>"
have "\<eta>\<^sub>1 (cod \<mu>) \<cdot> (PoE.map \<mu> \<star> e (src \<mu>))
= ((e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?g, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a] \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)"
unfolding unit\<^sub>1_def P_def emb.map_def hcomp_def src_def
- using \<mu> comp_char cod_simp arr_char P\<^sub>0_props [of ?a] P\<^sub>0_props [of ?b]
+ using \<mu> comp_char cod_simp arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props [of ?a] P\<^sub>0_props [of ?b]
by auto (* 10 sec *)
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?g, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)"
using 1 B.comp_assoc by simp
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?g, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a] \<cdot>\<^sub>B (e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a \<cdot>\<^sub>B e ?a)"
using 1 B.interchange [of "\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a]" "e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a" "e ?a" "e ?a"]
by simp
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?g, d ?a, e ?a] \<cdot>\<^sub>B
(((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a \<cdot>\<^sub>B e ?a)"
using 1 B.assoc'_naturality [of "e ?b" \<mu> "d ?a"] by simp
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?b \<star>\<^sub>B ?g, d ?a, e ?a] \<cdot>\<^sub>B
(((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)"
using 1 B.comp_assoc
B.interchange [of "(e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?a" "\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a]" "e ?a" "e ?a" ]
by simp
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
(((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)"
using 1 B.assoc_naturality [of "e ?b \<star>\<^sub>B \<mu>" "d ?a" "e ?a"] B.comp_assoc by simp
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)"
proof -
have "((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B ((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)
= (e ?b \<star>\<^sub>B ?g) \<cdot>\<^sub>B (e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B B.inv (\<eta> ?a) \<cdot>\<^sub>B (d ?a \<star>\<^sub>B e ?a)"
using 1 B.interchange by simp
also have "... = (e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B B.inv (\<eta> ?a)"
using 1 B.comp_arr_dom B.comp_cod_arr by simp
finally have "((e ?b \<star>\<^sub>B ?g) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B ((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)
= (e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B B.inv (\<eta> ?a)"
by simp
thus ?thesis by simp
qed
also have "... = (e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?b, ?g, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B ?a)) \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)"
proof -
have "(e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B B.inv (\<eta> ?a) = (e ?b \<star>\<^sub>B \<mu>) \<cdot>\<^sub>B (e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B ?a \<cdot>\<^sub>B B.inv (\<eta> ?a)"
using 1 B.comp_arr_dom B.comp_cod_arr by simp
also have "... = ((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B ?a) \<cdot>\<^sub>B ((e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B B.inv (\<eta> ?a))"
using 1 B.interchange by simp
finally have "(e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B B.inv (\<eta> ?a)
= ((e ?b \<star>\<^sub>B \<mu>) \<star>\<^sub>B ?a) \<cdot>\<^sub>B ((e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B B.inv (\<eta> ?a))"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B
(e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?f, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)"
using 1 B.assoc_naturality [of "e ?b" \<mu> "?a"] B.comp_assoc by simp
also have "... = (e ?b \<star>\<^sub>B \<mu>) \<cdot>\<^sub>B
(e ?b \<star>\<^sub>B \<r>\<^sub>B[?f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?f, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)"
proof -
have "(e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B (e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B ?a) = e ?b \<cdot>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[?g] \<cdot>\<^sub>B (\<mu> \<star>\<^sub>B ?a)"
using 1 B.interchange [of "e ?b" "e ?b" "\<r>\<^sub>B[?g]" "\<mu> \<star>\<^sub>B ?a"] by simp
also have "... = e ?b \<cdot>\<^sub>B e ?b \<star>\<^sub>B \<mu> \<cdot>\<^sub>B \<r>\<^sub>B[?f]"
using 1 B.runit_naturality by simp
also have "... = (e ?b \<star>\<^sub>B \<mu>) \<cdot>\<^sub>B (e ?b \<star>\<^sub>B \<r>\<^sub>B[?f])"
using 1 B.interchange [of "e ?b" "e ?b" \<mu> "\<r>\<^sub>B[?f]"] by simp
finally have "(e ?b \<star>\<^sub>B \<r>\<^sub>B[?g]) \<cdot>\<^sub>B (e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B ?a) = (e ?b \<star>\<^sub>B \<mu>) \<cdot>\<^sub>B (e ?b \<star>\<^sub>B \<r>\<^sub>B[?f])"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e (trg \<mu>) \<star> I\<^sub>S.map \<mu>) \<cdot> \<eta>\<^sub>1 (dom \<mu>)"
proof -
have "arr ((e (trg \<mu>) \<star> I\<^sub>S.map \<mu>) \<cdot> \<eta>\<^sub>1 (dom \<mu>))"
using \<mu> by simp
hence 2: "arr ((e ?b \<star>\<^sub>B \<mu>) \<cdot>
((e ?b \<star>\<^sub>B \<r>\<^sub>B[?f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, ?f, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B ?f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B ?f, d ?a, e ?a] \<cdot>\<^sub>B
(B.inv \<a>\<^sub>B[e ?b, ?f, d ?a] \<star>\<^sub>B e ?a)))"
unfolding unit\<^sub>1_def
- using \<mu> 1 arr_char P\<^sub>0_props trg_def hcomp_def dom_char by simp
+ using \<mu> 1 arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props trg_def hcomp_def dom_char\<^sub>S\<^sub>b\<^sub>C by simp
show ?thesis
unfolding unit\<^sub>1_def
- using \<mu> 1 arr_char P\<^sub>0_props trg_def hcomp_def dom_char
+ using \<mu> 1 arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props trg_def hcomp_def dom_char\<^sub>S\<^sub>b\<^sub>C
apply simp
using 2 comp_eqI by blast
qed
finally show "\<eta>\<^sub>1 (cod \<mu>) \<cdot> (PoE.map \<mu> \<star> e (src \<mu>))
= (e (trg \<mu>) \<star> I\<^sub>S.map \<mu>) \<cdot> \<eta>\<^sub>1 (dom \<mu>)"
by simp
qed
show "\<And>a. obj a \<Longrightarrow> (e a \<star> I\<^sub>S.unit a) \<cdot> \<r>\<^sup>-\<^sup>1[e a] \<cdot> \<l>[e a] = unit\<^sub>1 a \<cdot> (PoE.unit a \<star> e a)"
proof -
fix a
assume a: "obj a"
have 0: "B.obj a"
using a obj_char by blast
interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B \<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using 0 chosen_adjoint_equivalence(1) [of a] by simp
have src: "src\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) = P\<^sub>0 a"
using 0 B.src_vcomp [of "e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]" "B.inv (\<epsilon> a)"] emb.map_def by simp
have trg: "trg\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) = P\<^sub>0 a"
using 0 B.trg_vcomp [of "e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]" "B.inv (\<epsilon> a)"] emb.map_def by simp
have 1: "arr a \<and> src\<^sub>B a = a"
using a src_def by auto
have 2: "seq (P a) ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))"
proof -
have "B.arr (P a \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)))"
using a 0 P_def emb.map_def
by (elim B.objE, intro B.seqI) auto
moreover have "arr (P a)"
using 0 by auto
moreover have "arr ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))"
- apply (unfold arr_char)
+ apply (unfold arr_char\<^sub>S\<^sub>b\<^sub>C)
by (simp add: 0 P\<^sub>0_props(6) src trg)
ultimately show ?thesis
- using seq_char by simp
+ using seq_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
have 3: "obj (P\<^sub>0 a)"
- using 0 P\<^sub>0_props arr_char obj_char by blast
+ using 0 P\<^sub>0_props arr_char\<^sub>S\<^sub>b\<^sub>C obj_char by blast
have 4: "B.seq (P a) ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a))"
- using 2 seq_char by simp
+ using 2 seq_char\<^sub>S\<^sub>b\<^sub>C by simp
have "\<eta>\<^sub>1 a \<cdot> (PoE.unit a \<star> e a) = \<eta>\<^sub>1 a \<cdot> (P a \<cdot> (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
proof -
have "\<eta>\<^sub>1 a \<cdot> (PoE.unit a \<star> e a) = \<eta>\<^sub>1 a \<cdot> (PoE.unit a \<star>\<^sub>B e a)"
using 0 a hcomp_eqI by force
moreover have "PoE.unit a = P a \<cdot> (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)"
using a obj_char PoE.unit_char' emb.unit_char' prj_unit_char by simp
ultimately show ?thesis by argo
qed
also have "... = \<eta>\<^sub>1 a \<cdot> (P a \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
using 2 comp_simp by force
also have "... = \<eta>\<^sub>1 a \<cdot>\<^sub>B (P a \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
proof -
have "arr (\<eta>\<^sub>1 a)"
using a by auto
moreover have "arr (P a \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
by (metis (no_types, lifting) a 0 2 B.src_vcomp B.vseq_implies_hpar(1)
- arr_char arr_compI equivalence_data_in_hom(3) equivalence_data_simps\<^sub>B(6)
- hcomp_closed in_hom_char src)
+ arr_char\<^sub>S\<^sub>b\<^sub>C arr_compI equivalence_data_in_hom(3) equivalence_data_simps\<^sub>B(6)
+ hcomp_closed in_hom_char\<^sub>S\<^sub>b\<^sub>C src)
moreover have "B.seq (\<eta>\<^sub>1 a) (P a \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
using 0 3 P_def obj_char
by (elim B.objE, intro B.seqI) auto (* 10 sec *)
ultimately show ?thesis
using 2 comp_char by meson
qed
also have "... = (e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, e a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B e a) \<cdot>\<^sub>B
(((e a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
using 0 1 unit\<^sub>1_def P_def emb.map\<^sub>0_def emb.map_def B.comp_assoc B.obj_def' by simp
also have "... = (e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, e a] \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B e a) \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a))"
proof -
have "(e a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) = e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]"
using 0 B.comp_cod_arr by simp
thus ?thesis by simp
qed
also have "... = (e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B
((e a \<star>\<^sub>B a) \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B
(\<a>\<^sub>B[e a \<star>\<^sub>B a, d a, e a] \<cdot>\<^sub>B
((\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a) \<star>\<^sub>B e a)) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B e a) \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a)
= (\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B e a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B e a) \<cdot>\<^sub>B (B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
using 0 B.whisker_right by simp
also have "... = ((\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B e a) \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B e a)) \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
using B.comp_assoc by simp
also have "... = (\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<star>\<^sub>B e a) \<cdot>\<^sub>B (B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
proof -
have "B.seq \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using a 0 by (elim B.objE, intro B.seqI) auto
thus ?thesis
using 0 B.whisker_right by simp
qed
also have "... = ((\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a) \<star>\<^sub>B e a) \<cdot>\<^sub>B (B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
proof -
have "\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) = \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a"
proof -
have "e a \<star>\<^sub>B \<l>\<^sub>B[d a] = (\<r>\<^sub>B[e a] \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]"
using 0 1 B.triangle' [of "e a" "d a"] by simp
moreover have "B.hseq (e a) \<l>\<^sub>B[d a]"
using 0 by auto
moreover have "src\<^sub>B (e a) = a"
using 0 by simp
ultimately have "(\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a) \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B[d a]) = \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]"
using 0 B.invert_side_of_triangle(1)
[of "e a \<star>\<^sub>B \<l>\<^sub>B[d a]" "\<r>\<^sub>B[e a] \<star>\<^sub>B d a" "\<a>\<^sub>B\<^sup>-\<^sup>1[e a, src\<^sub>B (e a), d a]"]
by simp
moreover have "B.arr \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]"
using 0 B.assoc'_in_hom [of "e a" a "d a"] by fastforce
ultimately have "\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a = \<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<cdot>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using 0 B.invert_side_of_triangle(2)
[of "\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]" "\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a" "e a \<star>\<^sub>B \<l>\<^sub>B[d a]"]
by simp
thus ?thesis by simp
qed
thus ?thesis by simp
qed
finally have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \<star>\<^sub>B e a) \<cdot>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<star>\<^sub>B e a) =
((\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a) \<star>\<^sub>B e a) \<cdot>\<^sub>B (B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B
(((e a \<star>\<^sub>B a) \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B
(\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a \<star>\<^sub>B e a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, d a, e a] \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B e a)"
using 0 B.assoc_naturality [of "\<r>\<^sub>B\<^sup>-\<^sup>1[e a]" "d a" "e a"] B.comp_assoc by simp
also have "... = (e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B
(\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a) \<cdot>\<^sub>B
((e a \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, d a, e a] \<cdot>\<^sub>B
(B.inv (\<epsilon> a) \<star>\<^sub>B e a))"
proof -
have "((e a \<star>\<^sub>B a) \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a \<star>\<^sub>B e a)
= (e a \<star>\<^sub>B a) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B B.inv (\<eta> a) \<cdot>\<^sub>B (d a \<star>\<^sub>B e a)"
using 0 B.interchange [of "e a \<star>\<^sub>B a" "\<r>\<^sub>B\<^sup>-\<^sup>1[e a]" "B.inv (\<eta> a)" "d a \<star>\<^sub>B e a"]
by force
also have "... = \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B e a \<star>\<^sub>B a \<cdot>\<^sub>B B.inv (\<eta> a)"
using 0 B.comp_arr_dom B.comp_cod_arr by simp
also have "... = (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a) \<cdot>\<^sub>B (e a \<star>\<^sub>B B.inv (\<eta> a))"
using 0 B.interchange [of "\<r>\<^sub>B\<^sup>-\<^sup>1[e a]" "e a" a "B.inv (\<eta> a)"] by fastforce
ultimately have "((e a \<star>\<^sub>B a) \<star>\<^sub>B B.inv (\<eta> a)) \<cdot>\<^sub>B (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B d a \<star>\<^sub>B e a)
= (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a) \<cdot>\<^sub>B (e a \<star>\<^sub>B B.inv (\<eta> a))"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B
(\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a)) \<cdot>\<^sub>B
\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B \<l>\<^sub>B[e a]"
by (metis B.comp_assoc adjoint_equivalence_in_bicategory_def
adjunction_in_bicategory.triangle_right dual_adjoint_equivalence)
also have "... = \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B \<l>\<^sub>B[e a]"
proof -
have "(e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a) = e a \<star>\<^sub>B a"
proof -
have "(e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, a] \<cdot>\<^sub>B (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a)
= ((e a \<star>\<^sub>B \<r>\<^sub>B[a]) \<cdot>\<^sub>B \<a>\<^sub>B[e a, a, a]) \<cdot>\<^sub>B (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a)"
using B.comp_assoc by simp
also have "... = (\<r>\<^sub>B[e a] \<star>\<^sub>B a) \<cdot>\<^sub>B (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<star>\<^sub>B a)"
using 0 B.runit_char(2) [of "e a"] B.unitor_coincidence by simp
also have "... = e a \<star>\<^sub>B a"
using 0 B.whisker_right [of a "\<r>\<^sub>B[e a]" "\<r>\<^sub>B\<^sup>-\<^sup>1[e a]"] B.comp_arr_inv' by auto
finally show ?thesis
by simp
qed
thus ?thesis
using 0 B.comp_cod_arr by simp
qed
also have "... = (e a \<star> I\<^sub>S.unit a) \<cdot> \<r>\<^sup>-\<^sup>1[e a] \<cdot> \<l>[e a]"
proof -
have "(e a \<star> I\<^sub>S.unit a) \<cdot> \<r>\<^sup>-\<^sup>1[e a] \<cdot> \<l>[e a] = (e a \<star>\<^sub>B a) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B \<l>\<^sub>B[e a]"
proof -
have 1: "arr \<l>\<^sub>B[e a] \<and> arr \<r>\<^sub>B\<^sup>-\<^sup>1[e a]"
using a lunit_simp runit'_simp lunit_simps runit'_simps by simp
moreover have "arr (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B \<l>\<^sub>B[e a])"
- using 1 arr_char dom_char cod_char comp_char by auto
+ using 1 arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C comp_char by auto
moreover have 2: "hseq (e a) (I\<^sub>S.unit a)"
using a I\<^sub>S.unit_char by (intro hseqI') auto
moreover have "B.seq (e a \<star> I\<^sub>S.unit a) (\<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B \<l>\<^sub>B[e a])"
- using 0 2 hcomp_char arr_char I\<^sub>S.unit_char'
+ using 0 2 hcomp_char arr_char\<^sub>S\<^sub>b\<^sub>C I\<^sub>S.unit_char'
apply (intro B.seqI)
apply auto[4]
using a by force
ultimately show ?thesis
using a comp_char I\<^sub>S.unit_char' lunit_simp runit'_simp hcomp_char
by auto (* 4 sec *)
qed
also have "... = \<r>\<^sub>B\<^sup>-\<^sup>1[e a] \<cdot>\<^sub>B \<l>\<^sub>B[e a]"
using 0 B.comp_cod_arr by simp
finally show ?thesis by simp
qed
finally show "(e a \<star> I\<^sub>S.unit a) \<cdot> \<r>\<^sup>-\<^sup>1[e a] \<cdot> \<l>[e a] = \<eta>\<^sub>1 a \<cdot> (PoE.unit a \<star> e a)"
by argo
qed
show "\<And>f g. \<lbrakk>ide f; ide g; src g = trg f\<rbrakk> \<Longrightarrow>
(e (trg g) \<star> I\<^sub>S.cmp (g, f)) \<cdot>
\<a>[e (trg g), I\<^sub>S.map g, I\<^sub>S.map f] \<cdot>
(\<eta>\<^sub>1 g \<star> I\<^sub>S.map f) \<cdot>
inv \<a>[PoE.map g, e (src g), I\<^sub>S.map f] \<cdot>
(PoE.map g \<star> \<eta>\<^sub>1 f) \<cdot>
\<a>[PoE.map g, PoE.map f, e (src f)]
= \<eta>\<^sub>1 (g \<star> f) \<cdot> (PoE.cmp (g, f) \<star> e (src f))"
proof -
fix f g
assume f: "ide f" and g: "ide g" and fg: "src g = trg f"
let ?a = "src\<^sub>B f"
let ?b = "src\<^sub>B g"
let ?c = "trg\<^sub>B g"
have hseq: "ide f \<and> ide g \<and> arr f \<and> arr g \<and> src g = trg f"
- using f g fg ide_char by auto
+ using f g fg ide_char\<^sub>S\<^sub>b\<^sub>C by auto
have hseq\<^sub>B: "B.ide f \<and> B.ide g \<and> B.arr f \<and> B.arr g \<and> src\<^sub>B g = trg\<^sub>B f"
- using f g fg ide_char src_def trg_def by auto
+ using f g fg ide_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def by auto
have 1: "?a = src f \<and> ?b = src g \<and> ?c = trg g"
using f g fg src_def trg_def by simp
have "(e (trg g) \<star> I\<^sub>S.cmp (g, f)) \<cdot>
\<a>[e (trg g), I\<^sub>S.map g, I\<^sub>S.map f] \<cdot>
(\<eta>\<^sub>1 g \<star> I\<^sub>S.map f) \<cdot>
inv \<a>[PoE.map g, e (src g), I\<^sub>S.map f] \<cdot>
(PoE.map g \<star> \<eta>\<^sub>1 f) \<cdot>
\<a>[PoE.map g, PoE.map f, e (src f)]
= (e ?c \<star> I\<^sub>S.cmp (g, f)) \<cdot>
\<a>[e ?c, I\<^sub>S.map g, I\<^sub>S.map f] \<cdot>
(\<eta>\<^sub>1 g \<star> I\<^sub>S.map f) \<cdot>
inv \<a>[PoE.map g, e ?b, I\<^sub>S.map f] \<cdot>
(PoE.map g \<star> \<eta>\<^sub>1 f) \<cdot>
\<a>[PoE.map g, PoE.map f, e ?a]"
using 1 by simp
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(e ?c \<star> I\<^sub>S.cmp (g, f)) \<cdot>
\<a>[e ?c, I\<^sub>S.map g, I\<^sub>S.map f] \<cdot>
(\<eta>\<^sub>1 g \<star> I\<^sub>S.map f) \<cdot>
inv \<a>[PoE.map g, e ?b, I\<^sub>S.map f] \<cdot>
(PoE.map g \<star> unit\<^sub>1 f) \<cdot>
\<a>[PoE.map g, PoE.map f, e ?a]
= (e ?c \<star> g \<star> f) \<cdot>
\<a>[e ?c, g, f] \<cdot>
(\<eta>\<^sub>1 g \<star> f) \<cdot>
inv \<a>[P (E g), e ?b, f] \<cdot>
(P (E g) \<star> \<eta>\<^sub>1 f) \<cdot>
\<a>[P (E g), P (E f), e ?a]"
using f g fg by simp
also have "... = (e ?c \<star> g \<star> f) \<cdot>\<^sub>B
\<a>[e ?c, g, f] \<cdot>\<^sub>B
(\<eta>\<^sub>1 g \<star> f) \<cdot>\<^sub>B
inv \<a>[P (E g), e ?b, f] \<cdot>\<^sub>B
(P (E g) \<star> \<eta>\<^sub>1 f) \<cdot>\<^sub>B
\<a>[P (E g), P (E f), e ?a]"
proof -
have "arr ..."
proof -
have "\<guillemotleft>... : (P (E g) \<star> P (E f)) \<star> e ?a \<Rightarrow> e ?c \<star> g \<star> f\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>[P (E g), P (E f), e ?a] :
(P (E g) \<star> P (E f)) \<star> e ?a \<Rightarrow> P (E g) \<star> P (E f) \<star> e ?a\<guillemotright>"
- using f g fg VVV.arr_char VV.arr_char src_def obj_src assoc_in_hom
+ using f g fg VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def obj_src assoc_in_hom
by force
show "\<guillemotleft>P (E g) \<star> \<eta>\<^sub>1 f : P (E g) \<star> P (E f) \<star> e ?a \<Rightarrow> P (E g) \<star> e ?b \<star> f\<guillemotright>"
using f g fg 1 unit\<^sub>1_in_hom\<^sub>S [of f] vconn_implies_hpar(2)
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>inv \<a>[P (E g), e ?b, f] : P (E g) \<star> e ?b \<star> f \<Rightarrow> (P (E g) \<star> e ?b) \<star> f\<guillemotright>"
proof -
have "\<guillemotleft>\<a>[P (E g), e ?b, f] : (P (E g) \<star> e ?b) \<star> f \<Rightarrow> P (E g) \<star> e ?b \<star> f\<guillemotright>"
- using f g fg 1 VVV.arr_char VV.arr_char assoc_in_hom by simp
+ using f g fg 1 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C assoc_in_hom by simp
moreover have "iso \<a>[P (E g), e ?b, f]"
- using f g fg 1 VVV.arr_char VV.arr_char iso_assoc by simp
+ using f g fg 1 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C iso_assoc by simp
ultimately show ?thesis
using inv_in_hom by simp
qed
show "\<guillemotleft>unit\<^sub>1 g \<star> f : (P (E g) \<star> e ?b) \<star> f \<Rightarrow> (e ?c \<star> g) \<star> f\<guillemotright>"
using f g fg 1 unit\<^sub>1_in_hom\<^sub>S [of g]
by (intro hcomp_in_vhom) auto
have 2: "ide (e (trg g))"
using g by simp
have 3: "src (e (trg g)) = trg g"
- using g trg_def hseqE in_hom_char obj_trg by auto
+ using g trg_def hseqE in_hom_char\<^sub>S\<^sub>b\<^sub>C obj_trg by auto
show "\<guillemotleft>\<a>[e ?c, g, f] : (e ?c \<star> g) \<star> f \<Rightarrow> e ?c \<star> g \<star> f\<guillemotright>"
- using f g fg 1 2 3 VVV.arr_char VV.arr_char assoc_in_hom [of "e ?c" g f]
+ using f g fg 1 2 3 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C assoc_in_hom [of "e ?c" g f]
by simp
show "\<guillemotleft>e ?c \<star> g \<star> f : e ?c \<star> g \<star> f \<Rightarrow> e ?c \<star> g \<star> f\<guillemotright>"
using f g fg 1 2 3
by (intro hcomp_in_vhom) auto
qed
thus ?thesis by blast
qed
thus ?thesis
using comp_eqI by blast
qed
also have "... = (e ?c \<star> g \<star> f) \<cdot>\<^sub>B
\<a>[e ?c, g, f] \<cdot>\<^sub>B
(\<eta>\<^sub>1 g \<star> f) \<cdot>\<^sub>B
B.inv \<a>[P (E g), e ?b, f] \<cdot>\<^sub>B
(P (E g) \<star> \<eta>\<^sub>1 f) \<cdot>\<^sub>B
\<a>[P (E g), P (E f), e ?a]"
proof -
(*
* TODO: Maybe assoc should be a definition in subcategory, rather than
* an abbreviation. The expansion seems to create problems here.
*)
have "iso \<a>[P (E g), e ?b, f]"
- using f g fg hseq hseq\<^sub>B src_def trg_def emb.map\<^sub>0_def emb.map_def arr_char
+ using f g fg hseq hseq\<^sub>B src_def trg_def emb.map\<^sub>0_def emb.map_def arr_char\<^sub>S\<^sub>b\<^sub>C
obj_char P\<^sub>0_props(6)
by (intro iso_assoc) auto
thus ?thesis
- using f g fg inv_char [of "\<a>[P (E g), e ?b, f]"] by simp
+ using f g fg inv_char\<^sub>S\<^sub>b\<^sub>C [of "\<a>[P (E g), e ?b, f]"] by simp
qed
also have "... = (e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>[e ?c, g, f] \<cdot>\<^sub>B
(\<eta>\<^sub>1 g \<star>\<^sub>B f) \<cdot>\<^sub>B
B.inv \<a>[P (E g), e ?b, f] \<cdot>\<^sub>B
(P (E g) \<star>\<^sub>B \<eta>\<^sub>1 f) \<cdot>\<^sub>B
\<a>[P (E g), P (E f), e ?a]"
proof -
have "arr (unit\<^sub>0 ?c)"
using hseq obj_trg trg_def by auto
moreover have "arr (g \<star>\<^sub>B f)"
- using hseq hseq\<^sub>B arr_char by auto
+ using hseq hseq\<^sub>B arr_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show ?thesis
unfolding hcomp_def
using f g fg hseq\<^sub>B src_def trg_def P\<^sub>0_props(1) emb.map\<^sub>0_def emb.map_def
by simp
qed
also have "... = (e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
(\<eta>\<^sub>1 g \<star>\<^sub>B f) \<cdot>\<^sub>B
B.inv \<a>\<^sub>B[P (E g), e ?b, f] \<cdot>\<^sub>B
(P (E g) \<star>\<^sub>B \<eta>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[P (E g), P (E f), e ?a]"
proof -
have "arr (e ?a) \<and> arr (e ?b) \<and> arr (e ?c)"
using hseq hseq\<^sub>B src_def trg_def obj_char src.preserves_arr trg.preserves_arr
by auto
thus ?thesis
- using f g fg VVV.arr_char VV.arr_char src_def trg_def hseq\<^sub>B
+ using f g fg VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def hseq\<^sub>B
emb.map\<^sub>0_def emb.map_def
by simp
qed
also have "... = (e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
(\<eta>\<^sub>1 g \<star>\<^sub>B f) \<cdot>\<^sub>B
(B.inv \<a>\<^sub>B[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<eta>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a])"
using f g fg 1 emb.map_def by simp
also have "... = (e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
(\<eta>\<^sub>1 g \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<eta>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a])"
- using hseq\<^sub>B P.preserves_ide ide_char B.assoc'_eq_inv_assoc [of "P g" "e ?b" f]
+ using hseq\<^sub>B P.preserves_ide ide_char\<^sub>S\<^sub>b\<^sub>C B.assoc'_eq_inv_assoc [of "P g" "e ?b" f]
by simp
also have "... = (e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<r>\<^sub>B[g]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, ?b] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b)
\<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B (e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b, f, ?a] \<cdot>\<^sub>B
((e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
unfolding unit\<^sub>1_def
using hseq\<^sub>B B.comp_assoc by auto
also have "... = ((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, f]) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<r>\<^sub>B[g]) \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
- using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left B.whisker_right B.comp_assoc
+ using hseq\<^sub>B ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide B.whisker_left B.whisker_right B.comp_assoc
by auto
also have "... = \<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<r>\<^sub>B[g]) \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B[e ?c, g, ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f)) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<cdot>\<^sub>B \<a>\<^sub>B[e ?c, g, f] = \<a>\<^sub>B[e ?c, g, f]"
using hseq\<^sub>B B.comp_cod_arr by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = \<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<r>\<^sub>B[g]) \<star>\<^sub>B f) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f)) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(\<a>\<^sub>B[e ?c, g, ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B (((e ?c \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f)
= ((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B (\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f)"
proof -
have "(\<a>\<^sub>B[e ?c, g, ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B (((e ?c \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f)
= \<a>\<^sub>B[e ?c, g, ?b] \<cdot>\<^sub>B ((e ?c \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f"
using hseq\<^sub>B B.whisker_right by auto
also have "... = (e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b)) \<cdot>\<^sub>B \<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f"
using hseq\<^sub>B B.assoc_naturality [of "e ?c" g "B.inv (\<eta> ?b)"] by simp
also have "... = ((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f)"
using hseq\<^sub>B B.whisker_right by auto
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[e ?c, g, f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
using hseq\<^sub>B B.whisker_left B.whisker_right B.comp_assoc by simp
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
((P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
using hseq\<^sub>B B.comp_assoc
B.assoc_naturality [of "e ?c" "\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))" f]
by fastforce
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
((P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B (P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))
= (P g \<star>\<^sub>B e ?b \<star>\<^sub>B f \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B (P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a])"
proof -
have "(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a]) \<cdot>\<^sub>B (P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))
= P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, ?a] \<cdot>\<^sub>B ((e ?b \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))"
- using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left by auto
+ using hseq\<^sub>B ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide B.whisker_left by auto
also have "... = P g \<star>\<^sub>B (e ?b \<star>\<^sub>B f \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]"
using hseq\<^sub>B B.assoc_naturality [of "e ?b" f "B.inv (\<eta> ?a)"] by auto
also have "... = (P g \<star>\<^sub>B e ?b \<star>\<^sub>B f \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a])"
- using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left by auto
+ using hseq\<^sub>B ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide B.whisker_left by auto
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \<cdot>\<^sub>B
(P g \<star>\<^sub>B e ?b \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
- using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left B.comp_assoc by auto
+ using hseq\<^sub>B ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide B.whisker_left B.comp_assoc by auto
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
((P g \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
using hseq\<^sub>B B.comp_assoc
B.assoc'_naturality [of "P g" "e ?b" "\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"]
by simp
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
((P g \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
proof -
have "((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B f) \<cdot>\<^sub>B
((P g \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<cdot>\<^sub>B (P g \<star>\<^sub>B e ?b) \<star>\<^sub>B
(f \<cdot>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))"
proof -
have "B.seq (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) (P g \<star>\<^sub>B e ?b)"
using hseq\<^sub>B P.preserves_ide P_def src_def trg_def by auto
moreover have "B.seq f (\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))"
using hseq\<^sub>B by simp
ultimately show ?thesis
using B.interchange
[of "\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b" "P g \<star>\<^sub>B e ?b"
f "\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"]
by presburger
qed
also have "... = (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<cdot>\<^sub>B (e ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B e ?b)
\<star>\<^sub>B f \<cdot>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.whisker_right P_def by auto
also have "... = (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by simp
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
((\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B ((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b)
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= (\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a))"
proof -
have "(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<cdot>\<^sub>B (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b)
\<star>\<^sub>B f \<cdot>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.interchange by simp
also have "... = \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<cdot>\<^sub>B (\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b)
\<star>\<^sub>B (\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.comp_cod_arr B.comp_arr_dom by simp
also have "... = (\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a))"
using hseq\<^sub>B B.interchange by auto
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= \<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<cdot>\<^sub>B \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b]
\<star>\<^sub>B f \<cdot>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.interchange by simp
also have "... = \<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<cdot>\<^sub>B \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b]
\<star>\<^sub>B (\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.comp_cod_arr B.comp_arr_dom by simp
also have "... = (\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.interchange by auto
finally have "(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f) \<cdot>\<^sub>B (\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b]
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= (\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= \<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B ((e ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b]
\<star>\<^sub>B (\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a))"
using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by simp
also have "... = \<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.interchange by simp
also have "... = (\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using B.comp_assoc by simp
also have "... = (((e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.assoc_naturality [of "e ?c" "g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b"
"\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"]
by auto
also have "... = ((e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using B.comp_assoc by simp
finally have "\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f] \<cdot>\<^sub>B (\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b]
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))
= ((e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
have "(e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))))
= e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<cdot>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b)
\<star>\<^sub>B f \<cdot>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.whisker_left B.interchange by fastforce
also have "... = e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by auto
finally have "(e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))))
= e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
by simp
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, ?a] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
(* Working from the above expression to the common form, as in the previous part. *)
have "(e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, ?a] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)
= (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, ?a] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
using B.comp_assoc by simp
also have "... = ((e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B (g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.assoc_naturality [of "e ?c" "g \<star>\<^sub>B f" "B.inv (\<eta> ?a)"] B.comp_assoc
by simp
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_left B.comp_assoc by auto
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a)
= ((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_right B.whisker_left B.triangle' comp_assoc by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (((g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a))) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)
= (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a"
using hseq\<^sub>B B.whisker_right B.whisker_left by auto
also have "... = (e ?c \<star>\<^sub>B
(((g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])
\<star>\<^sub>B e ?a"
using hseq\<^sub>B B.assoc'_naturality [of g "B.inv (\<eta> ?b)" "f \<star>\<^sub>B d ?a"] by auto
also have "... = ((e ?c \<star>\<^sub>B (((g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a))) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_right B.whisker_left by auto
finally have "((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)
= ((e ?c \<star>\<^sub>B (((g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a))) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_right B.whisker_left B.comp_assoc by auto
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)
= (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a] \<cdot>\<^sub>B
((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a"
using hseq\<^sub>B B.whisker_left B.whisker_right by auto
also have "... = (e ?c \<star>\<^sub>B
(((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a])
\<star>\<^sub>B e ?a"
using hseq\<^sub>B B.assoc'_naturality [of "\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))" f "d ?a"]
by auto
also have "... = ((e ?c \<star>\<^sub>B (((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a))
\<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_left B.whisker_right by auto
finally have "((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)
= ((e ?c \<star>\<^sub>B (((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a)
= \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a"
using hseq\<^sub>B B.whisker_right by auto
also have "... = ((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a"
using hseq\<^sub>B B.assoc'_naturality [of "e ?c" "(\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f" "d ?a"]
by fastforce
also have "... = (((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a)
\<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_right by auto
finally have "(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B (((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a)) \<star>\<^sub>B e ?a)
= (((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a)
= ((e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f)) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a]"
using hseq\<^sub>B
B.assoc_naturality [of "(e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f))"
"d ?a" "e ?a"]
by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.comp_assoc
B.assoc_naturality [of "e ?c" "(\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f"
"d ?a \<star>\<^sub>B e ?a"]
by force
also have "... = (e ?c \<star>\<^sub>B (g \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B[g, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B ((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))) \<star>\<^sub>B f) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a))
= (e ?c \<star>\<^sub>B (g \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<a>\<^sub>B[g, f, d ?a \<star>\<^sub>B e ?a])"
using hseq\<^sub>B B.runit_hcomp B.whisker_left B.comp_assoc
B.assoc_naturality [of g f "B.inv (\<eta> ?a)"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e ?c \<star>\<^sub>B (g \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<a>\<^sub>B[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "(e ?c \<star>\<^sub>B \<a>\<^sub>B[g, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)
= e ?c \<star>\<^sub>B \<a>\<^sub>B[g, f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
((\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_left by auto
also have "... = e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a \<star>\<^sub>B e ?a]"
using hseq\<^sub>B B.assoc_naturality [of "\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))" f "d ?a \<star>\<^sub>B e ?a"]
by auto
also have "... = (e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<a>\<^sub>B[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a \<star>\<^sub>B e ?a])"
using hseq\<^sub>B B.whisker_left by auto
finally have "(e ?c \<star>\<^sub>B \<a>\<^sub>B[g, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f) \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)
= (e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<a>\<^sub>B[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a \<star>\<^sub>B e ?a])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B
\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B \<a>\<^sub>B[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "(e ?c \<star>\<^sub>B (g \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a))
= e ?c \<star>\<^sub>B (g \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
(\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_left by auto
also have "... = e ?c \<star>\<^sub>B g \<cdot>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))
\<star>\<^sub>B (\<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.interchange by auto
also have "... = e ?c \<star>\<^sub>B (g \<cdot>\<^sub>B \<r>\<^sub>B[g]) \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)"
using B.comp_assoc by simp
also have "... = e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by auto
finally have "(e ?c \<star>\<^sub>B (g \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a)))) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B (\<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a))
= e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b)) \<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))"
by simp
thus ?thesis
by simp
qed
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g] \<cdot>\<^sub>B (g \<star>\<^sub>B B.inv (\<eta> ?b))
\<star>\<^sub>B \<r>\<^sub>B[f] \<cdot>\<^sub>B (f \<star>\<^sub>B B.inv (\<eta> ?a))) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
proof -
let ?LHS = "(e ?c \<star>\<^sub>B \<a>\<^sub>B[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B (g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b)) \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a])) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
let ?RHS = "\<a>\<^sub>B[e ?c, g \<star>\<^sub>B d ?b \<star>\<^sub>B e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, e ?b] \<star>\<^sub>B f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B e ?b) \<star>\<^sub>B (f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \<star>\<^sub>B d ?a \<star>\<^sub>B e ?a] \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b, f, d ?a \<star>\<^sub>B e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B[e ?b \<star>\<^sub>B f, d ?a, e ?a]) \<cdot>\<^sub>B
(P g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[P g, P f, e ?a]"
have "?LHS = ?RHS"
proof -
let ?LHSt = "(\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>, \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>)) \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(((\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>])) \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>)"
let ?RHSt = "\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>\<langle>e ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>) \<^bold>\<star> ( \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>)) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>P\<^bold>R\<^bold>J g, \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J g \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J g \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>, \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>P\<^bold>R\<^bold>J g \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>e ?a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>P\<^bold>R\<^bold>J g, \<^bold>P\<^bold>R\<^bold>J f, \<^bold>\<langle>e ?a\<^bold>\<rangle>\<^bold>]"
have "?LHS = \<lbrace>?LHSt\<rbrace>"
- using f g fg hseq\<^sub>B B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def
+ using f g fg hseq\<^sub>B B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def
by auto
also have "... = \<lbrace>?RHSt\<rbrace>"
using hseq\<^sub>B by (intro EV.eval_eqI, auto)
also have "... = ?RHS"
- using f g fg hseq\<^sub>B B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def
+ using f g fg hseq\<^sub>B B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def
by auto
finally show ?thesis by blast
qed
thus ?thesis using hseq\<^sub>B by auto
qed
finally show ?thesis by simp
qed
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot> (PoE.cmp (g, f) \<star> e (src f))"
proof -
have "\<eta>\<^sub>1 (g \<star> f) \<cdot> (PoE.cmp (g, f) \<star> e (src f))
= \<eta>\<^sub>1 (g \<star> f) \<cdot> (P (E (g \<star> f)) \<cdot> P (\<Phi>\<^sub>E (g, f)) \<cdot> \<Phi>\<^sub>P (E g, E f) \<star> e (src f))"
- using f g fg PoE.cmp_def VV.arr_char VV.dom_char by simp
+ using f g fg PoE.cmp_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.dom_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot> (P (E (g \<star> f)) \<cdot> P (g \<star>\<^sub>B f) \<cdot> \<Phi>\<^sub>P (E g, E f) \<star> e (src f))"
- using f g fg emb.cmp_def VV.arr_char by simp
+ using f g fg emb.cmp_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot> (P (E (g \<star> f)) \<cdot> \<Phi>\<^sub>P (E g, E f) \<star> e (src f))"
using f g fg hseq\<^sub>B comp_cod_arr emb.map_def hseq_char' prj.cmp_simps'(1,5)
by auto
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot> (P (g \<star>\<^sub>B f) \<cdot> \<Phi>\<^sub>P (g, f) \<star> e (src f))"
using hseq\<^sub>B hseq emb.map_def hcomp_char hseqI' by auto
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot> (\<Phi>\<^sub>P (g, f) \<cdot> (P g \<star> P f) \<star> e (src f))"
- using hseq\<^sub>B B.VV.arr_char B.VV.dom_char B.VV.cod_char \<Phi>\<^sub>P.naturality [of "(g, f)"]
+ using hseq\<^sub>B B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>P.naturality [of "(g, f)"]
P.FF_def
by auto
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot> (\<Phi>\<^sub>P (g, f) \<star> e (src f))"
proof -
have "\<Phi>\<^sub>P (g, f) \<cdot> (P g \<star> P f) = \<Phi>\<^sub>P (g, f)"
- using f g fg comp_arr_dom [of "\<Phi>\<^sub>P (g, f)" "P g \<star> P f"] VV.arr_char
+ using f g fg comp_arr_dom [of "\<Phi>\<^sub>P (g, f)" "P g \<star> P f"] VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) prj.cmp_simps'(1) \<Phi>\<^sub>P_simps(4) hseq\<^sub>B)
thus ?thesis by simp
qed
also have "... = \<eta>\<^sub>1 (g \<star> f) \<cdot>\<^sub>B (\<Phi>\<^sub>P (g, f) \<star> e (src f))"
proof -
(* TODO: For some reason, this requires an epic struggle. *)
have "seq (\<eta>\<^sub>1 (g \<star> f)) (\<Phi>\<^sub>P (g, f) \<star> e (src f))"
proof -
have "cod (\<Phi>\<^sub>P (g, f)) = P (g \<star> f)"
- using f g fg \<Phi>\<^sub>P_simps(5) [of g f] VV.arr_char
+ using f g fg \<Phi>\<^sub>P_simps(5) [of g f] VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) hcomp_eqI hseq hseq\<^sub>B hseqI')
moreover have "hseq (\<Phi>\<^sub>P (g, f)) (e (src f))"
- using f g fg hseq\<^sub>B 1 B.VV.arr_char by auto
+ using f g fg hseq\<^sub>B 1 B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show ?thesis
- using f g fg hseq\<^sub>B 1 B.VV.arr_char P\<^sub>0_props prj.cmp_simps emb.map_def
- \<Phi>\<^sub>P_simps(5) [of g f] hcomp_eqI B.VV.dom_char B.VV.cod_char P.FF_def
+ using f g fg hseq\<^sub>B 1 B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props prj.cmp_simps emb.map_def
+ \<Phi>\<^sub>P_simps(5) [of g f] hcomp_eqI B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C P.FF_def
by auto
qed
thus ?thesis
using comp_eqI by simp
qed
also have "... = \<eta>\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>B (\<Phi>\<^sub>P (g, f) \<star>\<^sub>B e (src f))"
- using f g fg hseq\<^sub>B 1 B.VV.arr_char hcomp_eqI by simp
+ using f g fg hseq\<^sub>B 1 B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_eqI by simp
also have "... = ((e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, ?a] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(\<Phi>\<^sub>P (g, f) \<star>\<^sub>B e ?a)"
unfolding unit\<^sub>1_def
using hseq\<^sub>B 1 comp_char by simp
also have "... = ((e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, ?a] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
(((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B
\<l>\<^sub>B[f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f))
\<star>\<^sub>B e ?a)"
proof -
have "B.VV.ide (g, f)"
- using f g fg ide_char src_def trg_def B.VV.ide_char by auto
+ using f g fg ide_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C by auto
hence "\<Phi>\<^sub>P (g, f) = CMP g f"
unfolding \<Phi>\<^sub>P_def
- using f g fg ide_char CMP.map_simp_ide by simp
+ using f g fg ide_char\<^sub>S\<^sub>b\<^sub>C CMP.map_simp_ide by simp
also have "... = (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f)"
using hseq\<^sub>B CMP_def by auto
finally have "\<Phi>\<^sub>P (g, f) = (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f)"
by blast
thus ?thesis by simp
qed
also have "... = ((e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, src f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a)) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f)
\<star>\<^sub>B e ?a)"
using f g fg hseq\<^sub>B src_def trg_def B.whisker_left B.comp_assoc by simp
also have "... = (e ?c \<star>\<^sub>B \<r>\<^sub>B[g \<star>\<^sub>B f]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g \<star>\<^sub>B f, src f] \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B f) \<star>\<^sub>B B.inv (\<eta> ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g \<star>\<^sub>B f, d ?a, e ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g \<star>\<^sub>B f, d ?a] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
proof -
have "B.arr ((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f))"
- using hseq\<^sub>B ide_char P.preserves_ide P_def by auto
+ using hseq\<^sub>B ide_char\<^sub>S\<^sub>b\<^sub>C P.preserves_ide P_def by auto
hence "(e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f)
\<star>\<^sub>B e ?a
= ((e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
(\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<star>\<^sub>B e ?a) \<cdot>\<^sub>B
((\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<star>\<^sub>B e ?a)"
using hseq\<^sub>B B.whisker_right by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
finally show ?thesis
- using f ide_char src_def by simp
+ using f ide_char\<^sub>S\<^sub>b\<^sub>C src_def by simp
qed
finally show "(e (trg g) \<star> I\<^sub>S.cmp (g, f)) \<cdot>
\<a>[e (trg g), I\<^sub>S.map g, I\<^sub>S.map f] \<cdot>
(\<eta>\<^sub>1 g \<star> I\<^sub>S.map f) \<cdot>
inv \<a>[PoE.map g, e (src g), I\<^sub>S.map f] \<cdot>
(PoE.map g \<star> \<eta>\<^sub>1 f) \<cdot>
\<a>[PoE.map g, PoE.map f, e (src f)]
= \<eta>\<^sub>1 (g \<star> f) \<cdot> (PoE.cmp (g, f) \<star> e (src f))"
by blast
qed
qed
abbreviation (input) counit\<^sub>0
where "counit\<^sub>0 \<equiv> d"
definition counit\<^sub>1
where "counit\<^sub>1 f = \<a>\<^sub>B[d (trg\<^sub>B f), e (trg\<^sub>B f), f \<star>\<^sub>B d (src\<^sub>B f)] \<cdot>\<^sub>B
(\<eta> (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f)) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d (src\<^sub>B f)]"
abbreviation (input) \<epsilon>\<^sub>0
where "\<epsilon>\<^sub>0 \<equiv> counit\<^sub>0"
abbreviation (input) \<epsilon>\<^sub>1
where "\<epsilon>\<^sub>1 \<equiv> counit\<^sub>1"
lemma counit\<^sub>1_in_hom [intro]:
assumes "B.ide f"
shows "\<guillemotleft>\<epsilon>\<^sub>1 f : f \<star>\<^sub>B d (src\<^sub>B f) \<Rightarrow>\<^sub>B d (trg\<^sub>B f) \<star>\<^sub>B e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f)\<guillemotright>"
using assms B.iso_is_arr
by (unfold counit\<^sub>1_def, intro B.comp_in_homI' B.seqI B.hseqI' B.hcomp_in_vhom) auto
lemma counit\<^sub>1_simps [simp]:
assumes "B.ide f"
shows "B.arr (\<epsilon>\<^sub>1 f)"
and "src\<^sub>B (\<epsilon>\<^sub>1 f) = P\<^sub>0 (src\<^sub>B f)" and "trg\<^sub>B (\<epsilon>\<^sub>1 f) = trg\<^sub>B f"
and "B.dom (\<epsilon>\<^sub>1 f) = f \<star>\<^sub>B d (src\<^sub>B f)"
and "B.cod (\<epsilon>\<^sub>1 f) = d (trg\<^sub>B f) \<star>\<^sub>B e (trg\<^sub>B f) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f)"
and "B.iso (\<epsilon>\<^sub>1 f)"
using assms counit\<^sub>1_in_hom
apply auto
using B.vconn_implies_hpar(1)
apply fastforce
using B.vconn_implies_hpar(2)
apply fastforce
unfolding counit\<^sub>1_def
apply (intro B.isos_compose)
by auto
lemma technical:
assumes "B.ide f" and "B.ide g" and "src\<^sub>B g = trg\<^sub>B f"
shows "(\<epsilon>\<^sub>1 g \<star>\<^sub>B P f) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d (src\<^sub>B g), P f] \<cdot>\<^sub>B (g \<star>\<^sub>B \<epsilon>\<^sub>1 f)
= (\<a>\<^sub>B[d (trg\<^sub>B g), e (trg\<^sub>B g), g \<star>\<^sub>B d (src\<^sub>B g)] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d (trg\<^sub>B g) \<star>\<^sub>B e (trg\<^sub>B g), g, d (src\<^sub>B g)] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d (trg\<^sub>B g) \<star>\<^sub>B e (trg\<^sub>B g)) \<star>\<^sub>B g, d (src\<^sub>B g), P f] \<cdot>\<^sub>B
(((d (trg\<^sub>B g) \<star>\<^sub>B e (trg\<^sub>B g)) \<star>\<^sub>B g) \<star>\<^sub>B d (src\<^sub>B g) \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d (trg\<^sub>B g) \<star>\<^sub>B e (trg\<^sub>B g)) \<star>\<^sub>B g)
\<star>\<^sub>B \<a>\<^sub>B[d (src\<^sub>B g), e (src\<^sub>B g), f \<star>\<^sub>B d (src\<^sub>B f)]) \<cdot>\<^sub>B
((\<eta> (trg\<^sub>B g) \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B (\<eta> (src\<^sub>B g) \<star>\<^sub>B f \<star>\<^sub>B d (src\<^sub>B f))) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[src\<^sub>B g, f, d (src\<^sub>B f)] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d (src\<^sub>B f)))"
proof -
let ?a = "src\<^sub>B f"
let ?b = "src\<^sub>B g"
let ?c = "trg\<^sub>B g"
have "(\<epsilon>\<^sub>1 g \<star>\<^sub>B P f) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B (g \<star>\<^sub>B \<epsilon>\<^sub>1 f)
= (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b]
\<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
using assms counit\<^sub>1_def by simp
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
proof -
have "\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b]
\<star>\<^sub>B P f
= (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f)"
using assms B.iso_is_arr B.whisker_right P_def by auto
moreover have "g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a]
= (g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
using assms B.iso_is_arr B.whisker_left by simp
ultimately show ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[?c, g, d ?b] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
using assms B.lunit_hcomp [of g "d ?b"] B.comp_assoc by simp
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[?c, g, d ?b] \<star>\<^sub>B P f)) \<cdot>\<^sub>B
((\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
using assms B.whisker_right P_def B.comp_assoc by simp
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
proof -
have "((\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B (\<a>\<^sub>B[?c, g, d ?b] \<star>\<^sub>B P f) =
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<cdot>\<^sub>B \<a>\<^sub>B[?c, g, d ?b] \<star>\<^sub>B P f"
using assms B.whisker_right P_def B.iso_is_arr by simp
also have "... = \<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<cdot>\<^sub>B ((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b) \<star>\<^sub>B P f"
using assms B.assoc_naturality [of "\<eta> ?c" g "d ?b"] B.iso_is_arr by simp
also have "... = (\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b) \<star>\<^sub>B P f)"
using assms B.whisker_right P_def B.iso_is_arr by simp
finally have "((\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B (\<a>\<^sub>B[?c, g, d ?b] \<star>\<^sub>B P f) =
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b) \<star>\<^sub>B P f)"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
((((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[?c \<star>\<^sub>B g, d ?b, P f]) \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
proof -
have "((\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] =
\<a>\<^sub>B\<^sup>-\<^sup>1[?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b \<star>\<^sub>B P f)"
using assms B.assoc'_naturality [of "\<l>\<^sub>B\<^sup>-\<^sup>1[g]" "d ?b" "P f"] by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b \<star>\<^sub>B P f)) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
proof -
have "(((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b) \<star>\<^sub>B P f) \<cdot>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[?c \<star>\<^sub>B g, d ?b, P f] =
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B ((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f)"
using assms B.assoc'_naturality [of "\<eta> ?c \<star>\<^sub>B g" "d ?b" "P f"] B.iso_is_arr
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a])) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
using assms B.whisker_right B.iso_is_arr P_def B.comp_assoc by simp
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])"
proof -
have "((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a])
= (((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a])"
using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def
B.interchange [of "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]" g]
B.interchange [of "(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g" "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a]))"
proof -
have "((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B (g \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= (((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)"
using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def
B.interchange [of "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]" g]
B.interchange [of "(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g" "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a))"
proof -
have "((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])
= (\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a]"
using assms B.comp_arr_dom B.iso_is_arr P_def
B.interchange [of "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]" g "\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a" "\<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a]"]
by simp
also have "... = (\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B
(\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)"
using assms B.lunit_hcomp by simp
also have "... = ((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a))"
using assms B.comp_arr_dom B.iso_is_arr P_def
B.interchange [of "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]" g]
by simp
finally have "((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (g \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[f \<star>\<^sub>B d ?a])
= ((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a))"
by blast
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
sublocale counit: pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>E \<circ> P\<close> EoP.cmp I\<^sub>C.map I\<^sub>C.cmp counit\<^sub>0 counit\<^sub>1
proof
show "\<And>a. B.obj a \<Longrightarrow> B.ide (d a)"
by simp
show "\<And>a. B.obj a \<Longrightarrow> B.equivalence_map (d a)"
proof -
fix a
assume a: "B.obj a"
interpret Adj: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
\<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using a chosen_adjoint_equivalence by simp
show "B.equivalence_map (d a)"
using Adj.equivalence_in_bicategory_axioms B.equivalence_map_def Adj.dual_equivalence
by blast
qed
show "\<And>a. B.obj a \<Longrightarrow> \<guillemotleft>d a : src\<^sub>B ((E \<circ> P) a) \<rightarrow>\<^sub>B src\<^sub>B (I\<^sub>C.map a)\<guillemotright>"
unfolding emb.map_def by fastforce
show "\<And>f. B.ide f \<Longrightarrow> B.iso (\<epsilon>\<^sub>1 f)"
by simp
show "\<And>f. B.ide f \<Longrightarrow> \<guillemotleft>\<epsilon>\<^sub>1 f : I\<^sub>C.map f \<star>\<^sub>B d (src\<^sub>B f) \<Rightarrow>\<^sub>B d (trg\<^sub>B f) \<star>\<^sub>B (E \<circ> P) f\<guillemotright>"
unfolding counit\<^sub>1_def P_def emb.map_def
using P_def P_simps(1)
by (intro B.comp_in_homI' B.seqI B.hseqI') (auto simp add: B.iso_is_arr)
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> \<epsilon>\<^sub>1 (B.cod \<mu>) \<cdot>\<^sub>B (I\<^sub>C.map \<mu> \<star>\<^sub>B d (src\<^sub>B \<mu>))
= (d (trg\<^sub>B \<mu>) \<star>\<^sub>B (E \<circ> P) \<mu>) \<cdot>\<^sub>B \<epsilon>\<^sub>1 (B.dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "B.arr \<mu>"
let ?a = "src\<^sub>B \<mu>"
let ?b = "trg\<^sub>B \<mu>"
let ?f = "B.dom \<mu>"
let ?g = "B.cod \<mu>"
have "\<epsilon>\<^sub>1 ?g \<cdot>\<^sub>B (I\<^sub>C.map \<mu> \<star>\<^sub>B d ?a)
= (\<a>\<^sub>B[d ?b, e ?b, ?g \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B ?g \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?g \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(\<mu> \<star>\<^sub>B d ?a)"
- using \<mu> counit\<^sub>1_def P_def emb.map_def arr_char P\<^sub>0_props(1) by simp
+ using \<mu> counit\<^sub>1_def P_def emb.map_def arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) by simp
also have "... = (d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?b, e ?b, ?f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B ?f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?f \<star>\<^sub>B d ?a]"
proof -
have "(\<a>\<^sub>B[d ?b, e ?b, ?g \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B ?g \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?g \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(\<mu> \<star>\<^sub>B d ?a)
= \<a>\<^sub>B[d ?b, e ?b, ?g \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B ?g \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?g \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<mu> \<star>\<^sub>B d ?a)"
using B.comp_assoc by simp
also have "... = \<a>\<^sub>B[d ?b, e ?b, ?g \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((\<eta> ?b \<star>\<^sub>B ?g \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?f \<star>\<^sub>B d ?a]"
using \<mu> B.lunit'_naturality [of "\<mu> \<star>\<^sub>B d ?a"] B.comp_assoc by simp
also have "... = (\<a>\<^sub>B[d ?b, e ?b, ?g \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B ?f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?f \<star>\<^sub>B d ?a]"
proof -
have "(\<eta> ?b \<star>\<^sub>B ?g \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) =
\<eta> ?b \<cdot>\<^sub>B ?b \<star>\<^sub>B ?g \<cdot>\<^sub>B \<mu> \<star>\<^sub>B d ?a \<cdot>\<^sub>B d ?a"
using \<mu> B.iso_is_arr
B.interchange [of "\<eta> ?b" ?b "(?g \<star>\<^sub>B d ?a)" "\<mu> \<star>\<^sub>B d ?a"]
B.interchange [of "?g" \<mu> "d ?a" "d ?a"]
by simp
also have "... = (d ?b \<star>\<^sub>B e ?b) \<cdot>\<^sub>B \<eta> ?b \<star>\<^sub>B \<mu> \<cdot>\<^sub>B ?f \<star>\<^sub>B d ?a \<cdot>\<^sub>B d ?a"
using \<mu> B.comp_arr_dom B.comp_cod_arr apply simp
by (metis B.arrI equivalence_data_in_hom\<^sub>B(7) equivalence_data_simps\<^sub>B(17-18)
B.obj_trg)
also have "... = ((d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (\<eta> ?b \<star>\<^sub>B ?f \<star>\<^sub>B d ?a)"
using \<mu> B.iso_is_arr B.interchange [of "d ?b \<star>\<^sub>B e ?b" "\<eta> ?b" "\<mu> \<star>\<^sub>B d ?a" "?f \<star>\<^sub>B d ?a"]
B.interchange [of \<mu> "?f" "d ?a" "d ?a"]
by simp
finally have "(\<eta> ?b \<star>\<^sub>B ?g \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a)
= ((d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (\<eta> ?b \<star>\<^sub>B ?f \<star>\<^sub>B d ?a)"
by blast
thus ?thesis using B.comp_assoc by simp
qed
also have "... = (d ?b \<star>\<^sub>B e ?b \<star>\<^sub>B \<mu> \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?b, e ?b, ?f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?b \<star>\<^sub>B ?f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[?f \<star>\<^sub>B d ?a]"
using \<mu> B.assoc_naturality [of "d ?b" "e ?b" "\<mu> \<star>\<^sub>B d ?a"] B.comp_assoc by simp
finally show ?thesis by simp
qed
also have "... = (d ?b \<star>\<^sub>B (E \<circ> P) \<mu>) \<cdot>\<^sub>B \<epsilon>\<^sub>1 ?f"
- using \<mu> counit\<^sub>1_def P_def emb.map_def arr_char P\<^sub>0_props(1) P_simps\<^sub>B(1)
+ using \<mu> counit\<^sub>1_def P_def emb.map_def arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) P_simps\<^sub>B(1)
by (simp add: P\<^sub>0_props(6))
finally show "\<epsilon>\<^sub>1 (?g) \<cdot>\<^sub>B (I\<^sub>C.map \<mu> \<star>\<^sub>B d ?a) = (d ?b \<star>\<^sub>B (E \<circ> P) \<mu>) \<cdot>\<^sub>B \<epsilon>\<^sub>1 ?f"
by blast
qed
show "\<And>f g. \<lbrakk>B.ide f; B.ide g; src\<^sub>B g = trg\<^sub>B f\<rbrakk> \<Longrightarrow>
(d (trg\<^sub>B g) \<star>\<^sub>B EoP.cmp (g, f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d (trg\<^sub>B g), (E \<circ> P) g, (E \<circ> P) f] \<cdot>\<^sub>B
(\<epsilon>\<^sub>1 g \<star>\<^sub>B (E \<circ> P) f) \<cdot>\<^sub>B
B.inv \<a>\<^sub>B[I\<^sub>C.map g, d (src\<^sub>B g), (E \<circ> P) f] \<cdot>\<^sub>B (I\<^sub>C.map g \<star>\<^sub>B \<epsilon>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[I\<^sub>C.map g, I\<^sub>C.map f, d (src\<^sub>B f)]
= \<epsilon>\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>B (I\<^sub>C.cmp (g, f) \<star>\<^sub>B d (src\<^sub>B f))"
proof -
fix f g
assume f: "B.ide f" and g: "B.ide g" and fg: "src\<^sub>B g = trg\<^sub>B f"
let ?a = "src\<^sub>B f"
let ?b = "src\<^sub>B g"
let ?c = "trg\<^sub>B g"
have "(d ?c \<star>\<^sub>B EoP.cmp (g, f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, (E \<circ> P) g, (E \<circ> P) f] \<cdot>\<^sub>B
(\<epsilon>\<^sub>1 g \<star>\<^sub>B (E \<circ> P) f) \<cdot>\<^sub>B
B.inv \<a>\<^sub>B[I\<^sub>C.map g, d ?b, (E \<circ> P) f] \<cdot>\<^sub>B
(I\<^sub>C.map g \<star>\<^sub>B \<epsilon>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[I\<^sub>C.map g, I\<^sub>C.map f, d ?a]
= (d ?c \<star>\<^sub>B P (g \<star>\<^sub>B f) \<cdot>\<^sub>B CMP.map (g, f) \<cdot>\<^sub>B (P g \<star>\<^sub>B P f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
(\<epsilon>\<^sub>1 g \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B
(g \<star>\<^sub>B \<epsilon>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "src\<^sub>B (e ?c \<star>\<^sub>B g \<star>\<^sub>B d ?b) = trg\<^sub>B (e ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)"
- using f g fg src_def trg_def arr_char P_simps\<^sub>B(1)
- by (simp add: P\<^sub>0_props(1,6) arr_char)
+ using f g fg src_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C P_simps\<^sub>B(1)
+ by (simp add: P\<^sub>0_props(1,6) arr_char\<^sub>S\<^sub>b\<^sub>C)
moreover have "B.inv \<a>\<^sub>B[g, d ?b, P f] = \<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f]"
using f g fg by (simp add: P_def)
ultimately show ?thesis
using f g fg emb.map_def P.preserves_reflects_arr P.preserves_reflects_arr
P.preserves_reflects_arr \<Phi>\<^sub>P_def emb.cmp_def
- EoP.cmp_def B.VV.arr_char B.VV.dom_char VV.arr_char
+ EoP.cmp_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
qed
also have "... = (d ?c \<star>\<^sub>B CMP.map (g, f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
((\<epsilon>\<^sub>1 g \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \<cdot>\<^sub>B
(g \<star>\<^sub>B \<epsilon>\<^sub>1 f)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "\<guillemotleft>CMP.map (g, f) : P g \<star> P f \<Rightarrow> P (g \<star>\<^sub>B f)\<guillemotright>"
- using f g fg VV.arr_char VV.cod_char P.FF_def \<Phi>\<^sub>P_def \<Phi>\<^sub>P_in_hom(2) by auto
+ using f g fg VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C P.FF_def \<Phi>\<^sub>P_def \<Phi>\<^sub>P_in_hom(2) by auto
hence "\<guillemotleft>CMP.map (g, f) : P g \<star>\<^sub>B P f \<Rightarrow>\<^sub>B P (g \<star>\<^sub>B f)\<guillemotright>"
- using in_hom_char hcomp_char by auto
+ using in_hom_char\<^sub>S\<^sub>b\<^sub>C hcomp_char by auto
hence "P (g \<star>\<^sub>B f) \<cdot>\<^sub>B CMP.map (g, f) \<cdot>\<^sub>B (P g \<star>\<^sub>B P f) = CMP.map (g, f)"
using B.comp_arr_dom B.comp_cod_arr by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c \<star>\<^sub>B CMP.map (g, f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
using f g fg technical B.comp_assoc by simp
also have "... = (d ?c \<star>\<^sub>B (e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f] \<cdot>\<^sub>B
\<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
- using f g fg CMP.map_simp_ide CMP_def B.VV.ide_char B.VV.arr_char
+ using f g fg CMP.map_simp_ide CMP_def B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
B.whisker_left B.whisker_left B.comp_assoc
by simp
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
((d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a])) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
using f g fg B.whisker_left P_def B.comp_assoc by simp (* 11 sec *)
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
((d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a])
= \<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a]"
proof -
have "(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B[e ?c, g, d ?b \<star>\<^sub>B P f]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B[e ?c \<star>\<^sub>B g, d ?b, P f]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, P g, P f] \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, d ?b] \<star>\<^sub>B P f) \<cdot>\<^sub>B
\<a>\<^sub>B\<^sup>-\<^sup>1[(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g, d ?b, P f] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B d ?b \<star>\<^sub>B P f) \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B \<a>\<^sub>B[d ?b, e ?b, f \<star>\<^sub>B d ?a])
= \<lbrace>(\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J f\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J g, \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[(\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?c\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\<cdot>
(((\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?c\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\<cdot>
(((\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?c\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<star>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>d ?b\<^bold>\<rangle>, \<^bold>\<langle>e ?b\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>])\<rbrace>"
- using f g fg B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def B.\<ll>_ide_simp
+ using f g fg B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def B.\<ll>_ide_simp
by auto
also have "... = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle>, \<^bold>\<langle>e ?c\<^bold>\<rangle>,
\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>d ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?c\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>,
(\<^bold>\<langle>d ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>e ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = \<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a]"
- using f g fg B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def B.\<ll>_ide_simp
+ using f g fg B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def B.\<ll>_ide_simp
by auto
finally show ?thesis by blast
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a] =
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a)"
using f g fg
B.assoc_naturality
[of "d ?c" "e ?c" "g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, (d ?b \<star>\<^sub>B e ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a] =
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a)"
using f g fg
B.assoc_naturality
[of "d ?c \<star>\<^sub>B e ?c" g "B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= (\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"
proof -
have "(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<star>\<^sub>B B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= ((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g) \<cdot>\<^sub>B (\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B
(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)"
using f g fg
B.interchange [of "(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g" "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]"
"B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a"
"\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"]
by fastforce
also have "... = (\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B
(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)"
using f g fg B.comp_cod_arr [of "(\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g]" "(d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g"]
by (auto simp add: B.iso_is_arr)
also have "... = (\<eta> ?c \<star>\<^sub>B g) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"
proof -
have "(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= B.inv (\<eta> ?b) \<cdot>\<^sub>B \<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"
proof -
have "B.seq (B.inv (\<eta> ?b)) (\<eta> ?b)"
using f g fg chosen_adjoint_equivalence(5) by fastforce
thus ?thesis
using f g fg B.whisker_right by simp
qed
also have "... = ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"
using f g fg chosen_adjoint_equivalence(5) B.comp_inv_arr' by simp
finally have "(B.inv (\<eta> ?b) \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (\<eta> ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<a>\<^sub>B[d ?c \<star>\<^sub>B e ?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((\<eta> ?c \<star>\<^sub>B g) \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
using f g fg B.comp_assoc
B.whisker_right [of "?b \<star>\<^sub>B f \<star>\<^sub>B d ?a" "\<eta> ?c \<star>\<^sub>B g" "\<l>\<^sub>B\<^sup>-\<^sup>1[g]"]
by fastforce
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
((d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
using f g fg B.iso_is_arr B.comp_assoc
B.assoc_naturality [of "\<eta> ?c" g "?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"]
by simp
also have "... = (d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] =
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a])"
using f g fg B.iso_is_arr
B.assoc_naturality [of "d ?c" "e ?c" "g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B (\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= (\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a])"
using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr
B.interchange [of "d ?c \<star>\<^sub>B e ?c" "\<eta> ?c"
"g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]" "g \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a"]
B.interchange [of "\<eta> ?c" ?c
"g \<star>\<^sub>B f \<star>\<^sub>B d ?a" "g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = \<a>\<^sub>B[d ?c, e ?c, (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B f \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
(?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "(d ?c \<star>\<^sub>B e ?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, e ?c, g \<star>\<^sub>B f \<star>\<^sub>B d ?a] =
\<a>\<^sub>B[d ?c, e ?c, (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a])"
using f g fg B.assoc_naturality [of "d ?c" "e ?c" "\<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]"] by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = \<a>\<^sub>B[d ?c, e ?c, (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "((d ?c \<star>\<^sub>B e ?c) \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B (\<eta> ?c \<star>\<^sub>B g \<star>\<^sub>B f \<star>\<^sub>B d ?a)
= (\<eta> ?c \<star>\<^sub>B (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B (?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a])"
using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr
B.interchange [of "d ?c \<star>\<^sub>B e ?c" "\<eta> ?c" "\<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]" "g \<star>\<^sub>B f \<star>\<^sub>B d ?a"]
B.interchange [of "\<eta> ?c" ?c "(g \<star>\<^sub>B f) \<star>\<^sub>B d ?a" "\<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = \<epsilon>\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>B (I\<^sub>C.cmp (g, f) \<star>\<^sub>B d ?a)"
proof -
have "\<epsilon>\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>B (I\<^sub>C.cmp (g, f) \<star>\<^sub>B d ?a)
= \<a>\<^sub>B[d ?c, e ?c, (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
\<l>\<^sub>B\<^sup>-\<^sup>1[(g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B d ?a)"
using f g fg counit\<^sub>1_def B.comp_assoc by simp
also have "... = \<a>\<^sub>B[d ?c, e ?c, (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<eta> ?c \<star>\<^sub>B (g \<star>\<^sub>B f) \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "\<l>\<^sub>B\<^sup>-\<^sup>1[(g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B d ?a)
= (?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
proof -
have "\<l>\<^sub>B\<^sup>-\<^sup>1[(g \<star>\<^sub>B f) \<star>\<^sub>B d ?a] \<cdot>\<^sub>B ((g \<star>\<^sub>B f) \<star>\<^sub>B d ?a) =
\<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>)\<rbrace>"
using f g fg B.\<ll>_ide_simp by auto
also have "... = \<lbrace>(\<^bold>\<langle>?c\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>?c\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>?c\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>d ?a\<^bold>\<rangle>)) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>d ?a\<^bold>\<rangle>\<^bold>]\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = (?c \<star>\<^sub>B \<a>\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \<cdot>\<^sub>B
(?c \<star>\<^sub>B g \<star>\<^sub>B \<l>\<^sub>B[f \<star>\<^sub>B d ?a]) \<cdot>\<^sub>B
\<a>\<^sub>B[?c, g, ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[g] \<star>\<^sub>B ?b \<star>\<^sub>B f \<star>\<^sub>B d ?a) \<cdot>\<^sub>B
(g \<star>\<^sub>B \<a>\<^sub>B[?b, f, d ?a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[f] \<star>\<^sub>B d ?a)) \<cdot>\<^sub>B
\<a>\<^sub>B[g, f, d ?a]"
- using f g fg B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def B.\<ll>_ide_simp
+ using f g fg B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def B.\<ll>_ide_simp
by auto
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
finally show "(d ?c \<star>\<^sub>B EoP.cmp (g, f)) \<cdot>\<^sub>B
\<a>\<^sub>B[d ?c, (E \<circ> P) g, (E \<circ> P) f] \<cdot>\<^sub>B
(\<epsilon>\<^sub>1 g \<star>\<^sub>B (E \<circ> P) f) \<cdot>\<^sub>B
B.inv \<a>\<^sub>B[I\<^sub>C.map g, d ?b, (E \<circ> P) f] \<cdot>\<^sub>B (I\<^sub>C.map g \<star>\<^sub>B \<epsilon>\<^sub>1 f) \<cdot>\<^sub>B
\<a>\<^sub>B[I\<^sub>C.map g, I\<^sub>C.map f, d ?a]
= \<epsilon>\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>B (I\<^sub>C.cmp (g, f) \<star>\<^sub>B d ?a)"
by blast
qed
show "\<And>a. B.obj a \<Longrightarrow>
(d a \<star>\<^sub>B EoP.unit a) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a] = \<epsilon>\<^sub>1 a \<cdot>\<^sub>B (I\<^sub>C.unit a \<star>\<^sub>B d a)"
proof -
fix a
assume a: "B.obj a"
interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B \<open>e a\<close> \<open>d a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using a chosen_adjoint_equivalence by simp
have 0: "src\<^sub>B a = a \<and> trg\<^sub>B a = a"
using a by auto
have 1: "obj (P\<^sub>0 a)"
using a P\<^sub>0_props(1) by simp
have "(d a \<star>\<^sub>B EoP.unit a) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a]
= (d a \<star>\<^sub>B ((e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B src\<^sub>B (a \<star>\<^sub>B d (src\<^sub>B a))) \<cdot>\<^sub>B
\<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a]"
proof -
have "B.hseq (e (trg\<^sub>B a)) (a \<star>\<^sub>B d (src\<^sub>B a))"
using a by (elim B.objE, intro B.hseqI') auto
moreover have "arr (src\<^sub>B (d (src\<^sub>B a)))"
- using a arr_char P\<^sub>0_props(1) obj_char B.obj_simps(1) by auto
+ using a arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) obj_char B.obj_simps(1) by auto
moreover have "arr (src\<^sub>B (a \<star>\<^sub>B d (src\<^sub>B a)))"
- using a arr_char P\<^sub>0_props(1) obj_char calculation(1) by fastforce
+ using a arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) obj_char calculation(1) by fastforce
moreover have "src\<^sub>B (a \<star>\<^sub>B d (src\<^sub>B a)) \<in> Obj \<and> trg\<^sub>B (e (trg\<^sub>B a)) \<in> Obj"
- using a B.obj_simps P\<^sub>0_props obj_char arr_char by simp
+ using a B.obj_simps P\<^sub>0_props obj_char arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "P\<^sub>0 a \<cdot>\<^sub>B P\<^sub>0 a \<in> Obj"
- using 1 arr_char P\<^sub>0_props(1) obj_char
+ using 1 arr_char\<^sub>S\<^sub>b\<^sub>C P\<^sub>0_props(1) obj_char
by (metis (no_types, lifting) B.cod_trg B.obj_def' B.trg.as_nat_trans.is_natural_2)
moreover have "emb.unit (P\<^sub>0 (src\<^sub>B a)) = P\<^sub>0 (src\<^sub>B a)"
using a 0 1 emb.unit_char' P.map\<^sub>0_def src_def by simp
ultimately show ?thesis
using a emb.map_def EoP.unit_char' prj_unit_char emb.unit_char' P.map\<^sub>0_def P_def
- src_def arr_char obj_char
+ src_def arr_char\<^sub>S\<^sub>b\<^sub>C obj_char
by simp
qed
also have "... = (d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a) \<cdot>\<^sub>B P\<^sub>0 a) \<cdot>\<^sub>B
\<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a]"
using a 1 B.comp_assoc B.obj_simps by auto
also have "... = (d a \<star>\<^sub>B (e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a]"
using a B.comp_arr_dom by simp
also have "... = (d a \<star>\<^sub>B e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B (d a \<star>\<^sub>B B.inv (\<epsilon> a)) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a]"
using a B.whisker_left B.comp_assoc by simp
also have "... = ((d a \<star>\<^sub>B e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<a>\<^sub>B[d a, e a, d a]) \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B d a)"
using a triangle_right B.comp_assoc
B.invert_side_of_triangle(1) [of "\<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a]" "d a \<star>\<^sub>B \<epsilon> a"]
by simp
also have "... = \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B ((d a \<star>\<^sub>B e a) \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B d a)"
proof -
have "(d a \<star>\<^sub>B e a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]) \<cdot>\<^sub>B \<a>\<^sub>B[d a, e a, d a] =
\<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B ((d a \<star>\<^sub>B e a) \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using a B.assoc_naturality [of "d a" "e a" "\<l>\<^sub>B\<^sup>-\<^sup>1[d a]"] by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using a B.interchange [of "d a \<star>\<^sub>B e a" "\<eta> a" "\<l>\<^sub>B\<^sup>-\<^sup>1[d a]" "d a"]
B.comp_arr_dom B.comp_cod_arr
by simp
also have "... = \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B[a, a, d a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[a] \<star>\<^sub>B d a)"
proof -
have "\<a>\<^sub>B[a, a, d a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[a] \<star>\<^sub>B d a) = a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]"
proof -
(* TODO: I wanted to prove this directly, but missed some necessary trick. *)
have "\<a>\<^sub>B[a, a, d a] \<cdot>\<^sub>B (\<l>\<^sub>B\<^sup>-\<^sup>1[a] \<star>\<^sub>B d a) = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0\<^bold>] \<^bold>\<star> \<^bold>\<langle>d a\<^bold>\<rangle>)\<rbrace>"
- using a ide_char B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def B.\<ll>_ide_simp
+ using a ide_char B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def B.\<ll>_ide_simp
by auto
also have "... = \<lbrace>\<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>d a\<^bold>\<rangle>\<^bold>]\<rbrace>"
- using a ide_char by (intro EV.eval_eqI, auto)
+ using a ide_char\<^sub>S\<^sub>b\<^sub>C by (intro EV.eval_eqI, auto)
also have "... = a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a]"
- using a ide_char B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char B.VVV.arr_char
- B.VV.ide_char B.VV.arr_char P_def B.\<ll>_ide_simp
+ using a ide_char B.\<alpha>_def B.\<alpha>'.map_ide_simp B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C P_def B.\<ll>_ide_simp
by auto
finally show ?thesis by blast
qed
hence "\<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B[a, a, d a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[a] \<star>\<^sub>B d a)
= \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B (a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
by simp
also have "... = \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[d a])"
using a B.interchange [of "\<eta> a" a "a \<star>\<^sub>B d a" "\<l>\<^sub>B\<^sup>-\<^sup>1[d a]"] B.comp_arr_dom B.comp_cod_arr
by simp
finally show ?thesis by simp
qed
also have "... = \<epsilon>\<^sub>1 a \<cdot>\<^sub>B (I\<^sub>C.unit a \<star>\<^sub>B d a)"
proof -
have "\<epsilon>\<^sub>1 a \<cdot>\<^sub>B (I\<^sub>C.unit a \<star>\<^sub>B d a) =
\<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[a \<star>\<^sub>B d a] \<cdot>\<^sub>B (a \<star>\<^sub>B d a)"
using a 0 counit\<^sub>1_def I\<^sub>C.unit_char' B.comp_assoc by simp
also have "... = \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<l>\<^sub>B\<^sup>-\<^sup>1[a \<star>\<^sub>B d a]"
proof -
have "B.arr \<l>\<^sub>B\<^sup>-\<^sup>1[a \<star>\<^sub>B d a]"
using a
by (metis B.ide_hcomp B.lunit'_simps(1) B.objE equivalence_data_simps\<^sub>B(8) ide_right)
moreover have "B.dom \<l>\<^sub>B\<^sup>-\<^sup>1[a \<star>\<^sub>B d a] = a \<star>\<^sub>B d a"
using a
by (metis B.ide_hcomp B.lunit'_simps(4) B.objE antipar(1) equivalence_data_simps\<^sub>B(5)
ide_right)
ultimately show ?thesis
using a B.comp_arr_dom [of "\<l>\<^sub>B\<^sup>-\<^sup>1[a \<star>\<^sub>B d a]" "a \<star>\<^sub>B d a"] by simp
qed
also have "... = \<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B[a, a, d a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[a] \<star>\<^sub>B d a)"
using a B.lunit_hcomp(4) [of a "d a"] by auto
finally have "\<epsilon>\<^sub>1 a \<cdot>\<^sub>B (I\<^sub>C.unit a \<star>\<^sub>B d a) =
\<a>\<^sub>B[d a, e a, a \<star>\<^sub>B d a] \<cdot>\<^sub>B (\<eta> a \<star>\<^sub>B a \<star>\<^sub>B d a) \<cdot>\<^sub>B \<a>\<^sub>B[a, a, d a] \<cdot>\<^sub>B
(\<l>\<^sub>B\<^sup>-\<^sup>1[a] \<star>\<^sub>B d a)"
by simp
thus ?thesis by simp
qed
finally show "(d a \<star>\<^sub>B EoP.unit a) \<cdot>\<^sub>B \<r>\<^sub>B\<^sup>-\<^sup>1[d a] \<cdot>\<^sub>B \<l>\<^sub>B[d a] = \<epsilon>\<^sub>1 a \<cdot>\<^sub>B (I\<^sub>C.unit a \<star>\<^sub>B d a)"
by blast
qed
qed
interpretation equivalence_of_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \<a> \<i>\<^sub>B src trg
E \<Phi>\<^sub>E P \<Phi>\<^sub>P unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1
..
lemma induces_equivalence:
shows "equivalence_of_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \<a> \<i>\<^sub>B src trg
E \<Phi>\<^sub>E P \<Phi>\<^sub>P unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
..
end
subsection "Equivalence Pseudofunctors, Bijective on Objects"
text \<open>
Here we carry out the extension of an equivalence pseudofunctor \<open>F\<close> to an equivalence
of bicategories in the special case that the object map of \<open>F\<close> is bijective.
The bijectivity assumption simplifies the construction of the unit and counit of the
equivalence (the components at objects are in fact identities), as well as the proofs
of the associated coherence conditions.
\<close>
locale equivalence_pseudofunctor_bij_on_obj =
equivalence_pseudofunctor +
assumes bij_on_obj: "bij_betw map\<^sub>0 (Collect C.obj) (Collect D.obj)"
begin
abbreviation F\<^sub>0
where "F\<^sub>0 \<equiv> map\<^sub>0"
definition G\<^sub>0
where "G\<^sub>0 = inv_into (Collect C.obj) F\<^sub>0"
lemma G\<^sub>0_props:
shows "D.obj b \<Longrightarrow> C.obj (G\<^sub>0 b) \<and> F\<^sub>0 (G\<^sub>0 b) = b"
and "C.obj a \<Longrightarrow> D.obj (F\<^sub>0 a) \<and> G\<^sub>0 (F\<^sub>0 a) = a"
proof -
have surj: "F\<^sub>0 ` (Collect C.obj) = Collect D.obj"
using bij_on_obj bij_betw_imp_surj_on by blast
assume b: "D.obj b"
show "C.obj (G\<^sub>0 b) \<and> F\<^sub>0 (G\<^sub>0 b) = b"
proof
show "C.obj (G\<^sub>0 b)"
unfolding G\<^sub>0_def
using b by (metis inv_into_into mem_Collect_eq surj)
show "F\<^sub>0 (G\<^sub>0 b) = b"
unfolding G\<^sub>0_def
using b by (simp add: f_inv_into_f surj)
qed
next
have bij: "bij_betw G\<^sub>0 (Collect D.obj) (Collect C.obj)"
using bij_on_obj G\<^sub>0_def bij_betw_inv_into by auto
have surj: "G\<^sub>0 ` (Collect D.obj) = Collect C.obj"
using bij_on_obj G\<^sub>0_def
by (metis bij_betw_def bij_betw_inv_into)
assume a: "C.obj a"
show "D.obj (F\<^sub>0 a) \<and> G\<^sub>0 (F\<^sub>0 a) = a"
using a bij surj G\<^sub>0_def bij_betw_imp_inj_on bij_on_obj by force
qed
text \<open>
We extend \<open>G\<^sub>0\<close> to all arrows of \<open>D\<close> using chosen adjoint equivalences, which extend \<open>F\<close>,
between \<open>hom\<^sub>C (a, a')\<close> and \<open>hom\<^sub>D (F a, F a')\<close>. The use of adjoint equivalences restricts
choices that prevent us from validating the necessary coherence conditions.
\<close>
definition G
where "G \<nu> = (if D.arr \<nu> then
equivalence_pseudofunctor_at_hom.G\<^sub>1 V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F
(G\<^sub>0 (src\<^sub>D \<nu>)) (G\<^sub>0 (trg\<^sub>D \<nu>)) \<nu>
else C.null)"
lemma G_in_hom [intro]:
assumes "D.arr \<nu>"
shows "\<guillemotleft>G \<nu> : G\<^sub>0 (src\<^sub>D \<nu>) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D \<nu>)\<guillemotright>"
and "\<guillemotleft>G \<nu> : G (D.dom \<nu>) \<Rightarrow>\<^sub>C G (D.cod \<nu>)\<guillemotright>"
proof -
have 1: "src\<^sub>D \<nu> = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<and> trg\<^sub>D \<nu> = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))"
using assms G\<^sub>0_props by simp
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : G\<^sub>0 (src\<^sub>D \<nu>) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D \<nu>)\<guillemotright>\<close>
using assms C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))\<guillemotright>\<close>
using assms 1 D.in_hhom_def D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D \<nu>)\<close> \<open>G\<^sub>0 (trg\<^sub>D \<nu>)\<close>
using assms G\<^sub>0_props(1) by unfold_locales auto
have 2: "hom\<^sub>D.arr \<nu>"
- using assms 1 hom\<^sub>D.arr_char by simp
+ using assms 1 hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show "\<guillemotleft>G \<nu> : G (D.dom \<nu>) \<Rightarrow>\<^sub>C G (D.cod \<nu>)\<guillemotright>"
unfolding G_def
- using assms 2 hom\<^sub>C.arr_char hom\<^sub>C.dom_char hom\<^sub>C.cod_char hom\<^sub>D.dom_char hom\<^sub>D.cod_char
+ using assms 2 hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.cod_char\<^sub>S\<^sub>b\<^sub>C
Faa'.\<eta>\<epsilon>.F.preserves_arr Faa'.\<eta>\<epsilon>.F.preserves_dom Faa'.\<eta>\<epsilon>.F.preserves_cod
by (intro C.in_homI) auto
thus "C.in_hhom (G \<nu>) (G\<^sub>0 (src\<^sub>D \<nu>)) (G\<^sub>0 (trg\<^sub>D \<nu>))"
proof -
have "hom\<^sub>C.arr (Faa'.G\<^sub>1 \<nu>)"
using 2 by simp
thus ?thesis
using G_def assms hom\<^sub>C.arrE by presburger
qed
qed
lemma G_simps [simp]:
assumes "D.arr \<nu>"
shows "C.arr (G \<nu>)"
and "src\<^sub>C (G \<nu>) = G\<^sub>0 (src\<^sub>D \<nu>)" and "trg\<^sub>C (G \<nu>) = G\<^sub>0 (trg\<^sub>D \<nu>)"
and "C.dom (G \<nu>) = G (D.dom \<nu>)" and "C.cod (G \<nu>) = G (D.cod \<nu>)"
using assms G_in_hom by auto
interpretation G: "functor" V\<^sub>D V\<^sub>C G
proof
show "\<And>f. \<not> D.arr f \<Longrightarrow> G f = C.null"
unfolding G_def by simp
fix \<nu>
assume \<nu>: "D.arr \<nu>"
show "C.arr (G \<nu>)"
using \<nu> by simp
show "C.dom (G \<nu>) = G (D.dom \<nu>)"
using \<nu> by simp
show "C.cod (G \<nu>) = G (D.cod \<nu>)"
using \<nu> by simp
next
fix \<mu> \<nu>
assume \<mu>\<nu>: "D.seq \<mu> \<nu>"
have 1: "D.arr \<mu> \<and> D.arr \<nu> \<and> src\<^sub>D \<mu> = src\<^sub>D \<nu> \<and> trg\<^sub>D \<mu> = trg\<^sub>D \<nu> \<and>
src\<^sub>D \<mu> = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<mu>)) \<and> trg\<^sub>D \<mu> = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<mu>)) \<and>
src\<^sub>D \<nu> = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<and> trg\<^sub>D \<nu> = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))"
using \<mu>\<nu> G\<^sub>0_props D.vseq_implies_hpar by auto
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : G\<^sub>0 (src\<^sub>D \<nu>) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D \<nu>)\<guillemotright>\<close>
using 1 C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))\<guillemotright>\<close>
using 1 D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret hom\<^sub>D': subcategory V\<^sub>D
\<open>\<lambda>\<nu>'. D.arr \<nu>' \<and> src\<^sub>D \<nu>' = src\<^sub>D \<nu> \<and> trg\<^sub>D \<nu>' = trg\<^sub>D \<nu>\<close>
using 1 D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D \<nu>)\<close> \<open>G\<^sub>0 (trg\<^sub>D \<nu>)\<close>
using \<mu>\<nu> 1 G\<^sub>0_props(1) by unfold_locales auto
have 2: "hom\<^sub>D.seq \<mu> \<nu>"
- using \<mu>\<nu> 1 hom\<^sub>D.arr_char hom\<^sub>D.comp_def by fastforce
+ using \<mu>\<nu> 1 hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.comp_def by fastforce
have "G (hom\<^sub>D.comp \<mu> \<nu>) = hom\<^sub>C.comp (G \<mu>) (G \<nu>)"
unfolding G_def
- using 1 2 G\<^sub>0_props Faa'.G\<^sub>1_props hom\<^sub>D.arr_char
+ using 1 2 G\<^sub>0_props Faa'.G\<^sub>1_props hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C
Faa'.\<eta>\<epsilon>.F.as_nat_trans.preserves_comp_2
by auto
thus "G (\<mu> \<cdot>\<^sub>D \<nu>) = G \<mu> \<cdot>\<^sub>C G \<nu>"
using \<mu>\<nu> 1 2 G_def hom\<^sub>C.comp_simp hom\<^sub>D.comp_simp D.src_vcomp D.trg_vcomp
Faa'.\<eta>\<epsilon>.F.preserves_arr
by metis
qed
lemma functor_G:
shows "functor V\<^sub>D V\<^sub>C G"
..
interpretation G: faithful_functor V\<^sub>D V\<^sub>C G
proof
fix f f'
assume par: "D.par f f'"
assume eq: "G f = G f'"
have 1: "src\<^sub>D f = src\<^sub>D f' \<and> trg\<^sub>D f = trg\<^sub>D f'"
using par by (metis D.src_dom D.trg_dom)
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : G\<^sub>0 (src\<^sub>D f) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D f)\<guillemotright>\<close>
using par C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D f)) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D f))\<guillemotright>\<close>
using par D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret F: equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D f)\<close> \<open>G\<^sub>0 (trg\<^sub>D f)\<close>
using par G\<^sub>0_props(1) by unfold_locales auto
interpret F': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D f')\<close> \<open>G\<^sub>0 (trg\<^sub>D f')\<close>
using par G\<^sub>0_props(1) by unfold_locales auto
have 2: "hom\<^sub>D.par f f'"
- using par 1 hom\<^sub>D.arr_char hom\<^sub>D.dom_char hom\<^sub>D.cod_char G\<^sub>0_props by simp
+ using par 1 hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.cod_char\<^sub>S\<^sub>b\<^sub>C G\<^sub>0_props by simp
have "F.G\<^sub>1 f = F.G\<^sub>1 f'"
using par eq G_def G_simps(2-3) by metis
thus "f = f'"
using F.\<eta>\<epsilon>.F_is_faithful
by (simp add: 2 faithful_functor_axioms_def faithful_functor_def)
qed
interpretation FG: composite_functor V\<^sub>D V\<^sub>C V\<^sub>D G F ..
interpretation FG: faithful_functor V\<^sub>D V\<^sub>D "F o G"
using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
by blast
interpretation GF: composite_functor V\<^sub>C V\<^sub>D V\<^sub>C F G ..
interpretation GF: faithful_functor V\<^sub>C V\<^sub>C "G o F"
using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
by blast
interpretation G: weak_arrow_of_homs V\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C src\<^sub>C trg\<^sub>C G
proof
have *: "\<And>b. D.obj b \<Longrightarrow> C.isomorphic (G b) (src\<^sub>C (G b))"
proof -
fix b
assume b: "D.obj b"
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : G\<^sub>0 b \<rightarrow>\<^sub>C G\<^sub>0 b\<guillemotright>\<close>
using b C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 b) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 b)\<guillemotright>\<close>
using b D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 b\<close> \<open>G\<^sub>0 b\<close>
using b G\<^sub>0_props(1) by unfold_locales auto
have 1: "C.in_hhom (G\<^sub>0 b) (G\<^sub>0 b) (G\<^sub>0 b)"
using b G\<^sub>0_props by force
text \<open>
Using the unit constraints of \<open>F\<close> and the fact that \<open>F\<^sub>0 (G\<^sub>0 b) = b\<close>,
we obtain an isomorphism \<open>\<guillemotleft>?\<psi> : b \<Rightarrow>\<^sub>D F (G\<^sub>0 b)\<guillemotright>\<close>, which is also
an isomorphism in \<open>hom\<^sub>D\<close>.
\<close>
let ?\<psi> = "unit (G\<^sub>0 b)"
have \<psi>: "\<guillemotleft>?\<psi> : b \<Rightarrow>\<^sub>D F (G\<^sub>0 b)\<guillemotright> \<and> D.iso ?\<psi>"
using b G\<^sub>0_props unit_char by auto
have \<psi>_in_hhom: "D.in_hhom ?\<psi> b b"
using b \<psi> 1 D.src_dom D.trg_dom by fastforce
have 2: "hom\<^sub>D.arr ?\<psi> \<and> hom\<^sub>D.arr (D.inv ?\<psi>)"
- by (metis D.in_hhomE D.inv_in_hhom G\<^sub>0_props(1) \<psi> \<psi>_in_hhom hom\<^sub>D.arr_char)
+ by (metis D.in_hhomE D.inv_in_hhom G\<^sub>0_props(1) \<psi> \<psi>_in_hhom hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C)
have \<psi>': "hom\<^sub>D.in_hom ?\<psi> b (Faa'.F\<^sub>1 (G\<^sub>0 b))"
proof
show "hom\<^sub>D.arr ?\<psi>"
using 2 by simp
show "hom\<^sub>D.dom ?\<psi> = b"
- using 2 \<psi> hom\<^sub>D.dom_char by auto
+ using 2 \<psi> hom\<^sub>D.dom_char\<^sub>S\<^sub>b\<^sub>C by auto
show "hom\<^sub>D.cod ?\<psi> = Faa'.F\<^sub>1 (G\<^sub>0 b)"
proof -
have "\<guillemotleft>?\<psi> : b \<Rightarrow>\<^sub>D Faa'.F\<^sub>1 (G\<^sub>0 b)\<guillemotright> \<and> D.iso ?\<psi>"
unfolding Faa'.F\<^sub>1_def using b \<psi> 1 G\<^sub>0_props by auto
thus ?thesis
- using 2 \<psi> hom\<^sub>D.cod_char by auto
+ using 2 \<psi> hom\<^sub>D.cod_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
qed
text \<open>
Transposing \<open>?\<psi>\<close> via the adjunction \<open>\<eta>\<epsilon>\<close> yields an isomorphism from
\<open>G b\<close> to \<open>G\<^sub>0 b\<close> in \<open>hom\<^sub>C\<close>, hence in \<open>C\<close>.
\<close>
have **: "hom\<^sub>C.isomorphic (G b) (G\<^sub>0 b)"
proof -
have "hom\<^sub>C.in_hom (Faa'.\<eta>\<epsilon>.\<psi> (G\<^sub>0 b) ?\<psi>) (Faa'.G\<^sub>1 b) (G\<^sub>0 b)"
- by (metis "1" C.ide_src C.in_hhom_def Faa'.\<eta>\<epsilon>.\<psi>_in_hom \<psi>' hom\<^sub>C.ideI)
+ by (metis "1" C.ide_src C.in_hhom_def Faa'.\<eta>\<epsilon>.\<psi>_in_hom \<psi>' hom\<^sub>C.ideI\<^sub>S\<^sub>b\<^sub>C)
moreover have "hom\<^sub>C.iso (Faa'.\<eta>\<epsilon>.\<psi> (G\<^sub>0 b) ?\<psi>)"
proof (unfold Faa'.\<eta>\<epsilon>.\<psi>_def)
have "hom\<^sub>C.iso_in_hom (hom\<^sub>C.comp (Faa'.\<epsilon> (G\<^sub>0 b))
(Faa'.G\<^sub>1 ?\<psi>)) (Faa'.G\<^sub>1 b) (G\<^sub>0 b)"
proof (intro hom\<^sub>C.comp_iso_in_hom)
show "hom\<^sub>C.iso_in_hom (Faa'.G\<^sub>1 (unit (G\<^sub>0 b)))
(Faa'.G\<^sub>1 b) (Faa'.\<eta>\<epsilon>.FG.map (G\<^sub>0 b))"
- using \<psi> \<psi>' 1 2 hom\<^sub>D.iso_char Faa'.\<eta>\<epsilon>.F.preserves_iso hom\<^sub>D.iso_char
- Faa'.\<eta>\<epsilon>.F.preserves_iso hom\<^sub>C.in_hom_char hom\<^sub>C.arr_char
+ using \<psi> \<psi>' 1 2 hom\<^sub>D.iso_char\<^sub>S\<^sub>b\<^sub>C Faa'.\<eta>\<epsilon>.F.preserves_iso hom\<^sub>D.iso_char\<^sub>S\<^sub>b\<^sub>C
+ Faa'.\<eta>\<epsilon>.F.preserves_iso hom\<^sub>C.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
show "hom\<^sub>C.iso_in_hom (Faa'.\<epsilon> (G\<^sub>0 b)) (Faa'.\<eta>\<epsilon>.FG.map (G\<^sub>0 b)) (G\<^sub>0 b)"
using 1 C.ide_src C.ide_trg C.in_hhom_def Faa'.\<eta>\<epsilon>.\<epsilon>.components_are_iso
- Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_hom hom\<^sub>C.arrI hom\<^sub>C.ideI hom\<^sub>C.ide_in_hom
+ Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_hom hom\<^sub>C.arrI hom\<^sub>C.ideI\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.ide_in_hom
hom\<^sub>C.map_simp hom\<^sub>C.iso_in_hom_def
by metis
qed
thus "hom\<^sub>C.iso (hom\<^sub>C.comp (Faa'.\<epsilon> (G\<^sub>0 b)) (Faa'.G\<^sub>1 (unit (G\<^sub>0 b))))" by auto
qed
ultimately show ?thesis
unfolding G_def
using b hom\<^sub>C.isomorphic_def D.obj_def D.obj_simps(3) by auto
qed
hence "hom\<^sub>C.isomorphic (G b) (src\<^sub>C (G b))"
using b G_in_hom(1) by auto
moreover have "\<And>f g. hom\<^sub>C.isomorphic f g \<Longrightarrow> C.isomorphic f g"
- using hom\<^sub>C.iso_char hom\<^sub>C.isomorphic_def C.isomorphic_def hom\<^sub>C.in_hom_char hom\<^sub>C.arr_char
+ using hom\<^sub>C.iso_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.isomorphic_def C.isomorphic_def hom\<^sub>C.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
ultimately show "C.isomorphic (G b) (src\<^sub>C (G b))"
by simp
qed
fix \<nu>
assume \<nu>: "D.arr \<nu>"
show "C.isomorphic (G (src\<^sub>D \<nu>)) (src\<^sub>C (G \<nu>))"
using \<nu> * by force
show "C.isomorphic (G (trg\<^sub>D \<nu>)) (trg\<^sub>C (G \<nu>))"
using \<nu> * by force
qed
lemma weak_arrow_of_homs_G:
shows "weak_arrow_of_homs V\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C src\<^sub>C trg\<^sub>C G"
..
sublocale H\<^sub>DoGG: composite_functor D.VV.comp C.VV.comp V\<^sub>C
G.FF \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>C snd \<mu>\<nu>\<close>
..
sublocale GoH\<^sub>D: composite_functor D.VV.comp V\<^sub>D V\<^sub>C \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D snd \<mu>\<nu>\<close> G
..
text \<open>
To get the unit \<open>\<eta>\<close> of the equivalence of bicategories, we piece together the
units from the local equivalences. The components at objects will in fact be identities.
\<close>
definition \<eta>
where "\<eta> \<nu> = (if D.arr \<nu> then
equivalence_pseudofunctor_at_hom.\<eta> V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F
(G\<^sub>0 (src\<^sub>D \<nu>)) (G\<^sub>0 (trg\<^sub>D \<nu>)) \<nu>
else D.null)"
lemma \<eta>_in_hom:
assumes "D.arr \<nu>"
shows [intro]: "\<guillemotleft>\<eta> \<nu> : src\<^sub>D \<nu> \<rightarrow>\<^sub>D trg\<^sub>D \<nu>\<guillemotright>"
and [intro]: "\<guillemotleft>\<eta> \<nu> : D.dom \<nu> \<Rightarrow>\<^sub>D F (G (D.cod \<nu>))\<guillemotright>"
and "D.ide \<nu> \<Longrightarrow> D.iso (\<eta> \<nu>)"
proof -
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : G\<^sub>0 (src\<^sub>D \<nu>) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D \<nu>)\<guillemotright>\<close>
using assms C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))\<guillemotright>\<close>
using assms D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D \<nu>)\<close> \<open>G\<^sub>0 (trg\<^sub>D \<nu>)\<close>
using assms G\<^sub>0_props(1) by unfold_locales auto
have 1: "\<guillemotleft>\<eta> \<nu> : D.dom \<nu> \<Rightarrow>\<^sub>D F (G (D.cod \<nu>))\<guillemotright> \<and> (D.ide \<nu> \<longrightarrow> D.iso (\<eta> \<nu>))"
proof -
have "src\<^sub>D \<nu> = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<and> trg\<^sub>D \<nu> = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))"
using assms G\<^sub>0_props by simp
hence "hom\<^sub>D.arr \<nu>"
- using assms hom\<^sub>D.arr_char hom\<^sub>D.ide_char by simp
+ using assms hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.ide_char by simp
hence "hom\<^sub>D.in_hom (Faa'.\<eta> \<nu>) (D.dom \<nu>) (F (G (D.cod \<nu>))) \<and>
(D.ide \<nu> \<longrightarrow> hom\<^sub>D.iso (Faa'.\<eta> \<nu>))"
using assms Faa'.\<eta>\<epsilon>.\<eta>.preserves_hom [of \<nu> "D.dom \<nu>" "D.cod \<nu>"]
hom\<^sub>D.map_def G_def Faa'.F\<^sub>1_def
apply simp
- by (simp add: D.arr_iff_in_hom hom\<^sub>D.arr_char hom\<^sub>D.cod_closed hom\<^sub>D.dom_closed
- hom\<^sub>D.ide_char hom\<^sub>D.in_hom_char)
+ by (simp add: D.arr_iff_in_hom hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.cod_closed hom\<^sub>D.dom_closed
+ hom\<^sub>D.ide_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.in_hom_char\<^sub>S\<^sub>b\<^sub>C)
thus ?thesis
unfolding \<eta>_def
- using hom\<^sub>D.in_hom_char hom\<^sub>D.iso_char by auto
+ using hom\<^sub>D.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.iso_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
show "\<guillemotleft>\<eta> \<nu> : D.dom \<nu> \<Rightarrow>\<^sub>D F (G (D.cod \<nu>))\<guillemotright>"
using 1 by simp
thus "D.in_hhom (\<eta> \<nu>) (src\<^sub>D \<nu>) (trg\<^sub>D \<nu>)"
using assms D.src_dom D.trg_dom
by (metis D.arrI D.in_hhom_def D.vconn_implies_hpar(1-2))
show "D.ide \<nu> \<Longrightarrow> D.iso (\<eta> \<nu>)"
using 1 by simp
qed
lemma \<eta>_simps [simp]:
assumes "D.arr \<nu>"
shows "D.arr (\<eta> \<nu>)"
and "src\<^sub>D (\<eta> \<nu>) = src\<^sub>D \<nu>" and "trg\<^sub>D (\<eta> \<nu>) = trg\<^sub>D \<nu>"
and "D.dom (\<eta> \<nu>) = D.dom \<nu>" and "D.cod (\<eta> \<nu>) = F (G (D.cod \<nu>))"
and "D.ide \<nu> \<Longrightarrow> D.iso (\<eta> \<nu>)"
using assms \<eta>_in_hom by auto
lemma \<eta>_naturality:
assumes "D.arr \<nu>"
shows "\<eta> (D.cod \<nu>) \<cdot>\<^sub>D \<nu> = \<eta> \<nu>" and "F (G \<nu>) \<cdot>\<^sub>D \<eta> (D.dom \<nu>) = \<eta> \<nu>"
and "\<nu> \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>)) = D.inv (\<eta> (D.cod \<nu>)) \<cdot>\<^sub>D F (G \<nu>)"
proof -
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>'. \<guillemotleft>\<mu>' : G\<^sub>0 (src\<^sub>D \<nu>) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D \<nu>)\<guillemotright>\<close>
using assms C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D \<nu>)\<close> \<open>G\<^sub>0 (trg\<^sub>D \<nu>)\<close>
using assms G\<^sub>0_props by unfold_locales auto
show 1: "\<eta> (D.cod \<nu>) \<cdot>\<^sub>D \<nu> = \<eta> \<nu>"
unfolding \<eta>_def
using assms Faa'.\<eta>\<epsilon>.\<eta>.is_natural_2 hom\<^sub>D.comp_char G_def hom\<^sub>D.cod_simp
- G_in_hom(1) hom\<^sub>C.arr_char hom\<^sub>D.arr_char hom\<^sub>D.cod_closed
+ G_in_hom(1) hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.cod_closed
apply simp
by (metis (no_types, lifting) D.ext Faa'.\<eta>\<epsilon>.F.preserves_reflects_arr
Faa'.\<eta>\<epsilon>.\<eta>.preserves_reflects_arr)
show 2: "F (G \<nu>) \<cdot>\<^sub>D \<eta> (D.dom \<nu>) = \<eta> \<nu>"
unfolding \<eta>_def
using assms D.src_dom D.src_cod D.trg_dom D.trg_cod Faa'.\<eta>\<epsilon>.\<eta>.is_natural_1
hom\<^sub>D.comp_char G_def Faa'.F\<^sub>1_def hom\<^sub>D.dom_simp hom\<^sub>D.cod_simp
apply simp
by (metis (no_types, lifting) D.not_arr_null Faa'.\<eta>\<epsilon>.\<eta>.is_extensional
\<eta>_def \<eta>_simps(1) hom\<^sub>D.null_char)
show "\<nu> \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>)) = D.inv (\<eta> (D.cod \<nu>)) \<cdot>\<^sub>D F (G \<nu>)"
using assms 1 2
by (simp add: D.invert_opposite_sides_of_square)
qed
text \<open>
The fact that \<open>G\<^sub>0\<close> is chosen to be right-inverse to \<open>F\<close> implies that \<open>\<eta>\<close> is an
ordinary natural isomorphism (with respect to vertical composition) from \<open>I\<^sub>D\<close> to \<open>FG\<close>.
\<close>
interpretation \<eta>: natural_isomorphism V\<^sub>D V\<^sub>D D.map FG.map \<eta>
using \<eta>_simps \<eta>_naturality \<eta>_def
by unfold_locales auto
text \<open>
In view of the bijectivity assumption, we can obtain the counit \<open>\<epsilon>\<close> the same way.
\<close>
definition \<epsilon>
where "\<epsilon> \<mu> = (if C.arr \<mu> then
equivalence_pseudofunctor_at_hom.\<epsilon> V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F
(src\<^sub>C \<mu>) (trg\<^sub>C \<mu>) \<mu>
else C.null)"
lemma \<epsilon>_in_hom:
assumes "C.arr \<mu>"
shows [intro]: "\<guillemotleft>\<epsilon> \<mu> : src\<^sub>C \<mu> \<rightarrow>\<^sub>C trg\<^sub>C \<mu>\<guillemotright>"
and [intro]: "\<guillemotleft>\<epsilon> \<mu> : G (F (C.dom \<mu>)) \<Rightarrow>\<^sub>C C.cod \<mu>\<guillemotright>"
and "C.ide \<mu> \<Longrightarrow> C.iso (\<epsilon> \<mu>)"
proof -
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>'. \<guillemotleft>\<mu>' : src\<^sub>C \<mu> \<rightarrow>\<^sub>C trg\<^sub>C \<mu>\<guillemotright>\<close>
using C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (src\<^sub>C \<mu>) \<rightarrow>\<^sub>D F\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>\<close>
using assms D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (F\<^sub>0 (src\<^sub>C \<mu>))\<close> \<open>G\<^sub>0 (F\<^sub>0 (trg\<^sub>C \<mu>))\<close>
using assms G\<^sub>0_props
by unfold_locales simp_all
have 1: "\<guillemotleft>\<epsilon> \<mu> : G (F (C.dom \<mu>)) \<Rightarrow>\<^sub>C C.cod \<mu>\<guillemotright> \<and> (C.ide \<mu> \<longrightarrow> C.iso (\<epsilon> \<mu>))"
proof -
have "hom\<^sub>C.arr \<mu>"
- using assms hom\<^sub>C.arr_char by simp
+ using assms hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
hence "hom\<^sub>C.in_hom \<mu> (C.dom \<mu>) (C.cod \<mu>)"
- using assms hom\<^sub>C.in_hom_char hom\<^sub>C.dom_char hom\<^sub>C.cod_char by auto
+ using assms hom\<^sub>C.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C by auto
hence "hom\<^sub>C.in_hom (Faa'.\<epsilon> \<mu>) (G (F (C.dom \<mu>))) (C.cod \<mu>) \<and>
(C.ide \<mu> \<longrightarrow> hom\<^sub>C.iso (Faa'.\<epsilon> \<mu>))"
using assms Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_hom [of \<mu> "C.dom \<mu>" "C.cod \<mu>"]
Faa'.\<eta>\<epsilon>.\<epsilon>.components_are_iso hom\<^sub>C.map_def G_def Faa'.F\<^sub>1_def G\<^sub>0_props
- by (simp add: hom\<^sub>C.ideI)
+ by (simp add: hom\<^sub>C.ideI\<^sub>S\<^sub>b\<^sub>C)
thus ?thesis
unfolding \<epsilon>_def
- using assms hom\<^sub>C.in_hom_char hom\<^sub>C.iso_char G\<^sub>0_props(2) by simp
+ using assms hom\<^sub>C.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.iso_char\<^sub>S\<^sub>b\<^sub>C G\<^sub>0_props(2) by simp
qed
show "\<guillemotleft>\<epsilon> \<mu> : G (F (C.dom \<mu>)) \<Rightarrow>\<^sub>C C.cod \<mu>\<guillemotright>"
using 1 by simp
thus "C.in_hhom (\<epsilon> \<mu>) (src\<^sub>C \<mu>) (trg\<^sub>C \<mu>)"
using assms C.src_cod C.trg_cod
by (metis C.arrI C.in_hhom_def C.vconn_implies_hpar(1-4))
show "C.ide \<mu> \<Longrightarrow> C.iso (\<epsilon> \<mu>)"
using 1 by simp
qed
lemma \<epsilon>_simps [simp]:
assumes "C.arr \<mu>"
shows "C.arr (\<epsilon> \<mu>)"
and "src\<^sub>C (\<epsilon> \<mu>) = src\<^sub>C \<mu>" and "trg\<^sub>C (\<epsilon> \<mu>) = trg\<^sub>C \<mu>"
and "C.dom (\<epsilon> \<mu>) = G (F (C.dom \<mu>))" and "C.cod (\<epsilon> \<mu>) = C.cod \<mu>"
and "C.ide \<mu> \<Longrightarrow> C.iso (\<epsilon> \<mu>)"
using assms \<epsilon>_in_hom by auto
lemma \<epsilon>_naturality:
assumes "C.arr \<mu>"
shows "\<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C G (F \<mu>) = \<epsilon> \<mu>" and "\<mu> \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) = \<epsilon> \<mu>"
and "G (F \<mu>) \<cdot>\<^sub>C C.inv (\<epsilon> (C.dom \<mu>)) = C.inv (\<epsilon> (C.cod \<mu>)) \<cdot>\<^sub>C \<mu>"
proof -
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>'. \<guillemotleft>\<mu>' : src\<^sub>C \<mu> \<rightarrow>\<^sub>C trg\<^sub>C \<mu>\<guillemotright>\<close>
using assms C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>. \<guillemotleft>\<nu> : F\<^sub>0 (src\<^sub>C \<mu>) \<rightarrow>\<^sub>D F\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>\<close>
using assms D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>src\<^sub>C \<mu>\<close> \<open>trg\<^sub>C \<mu>\<close>
using assms by unfold_locales simp_all
show 1: "\<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C G (F \<mu>) = \<epsilon> \<mu>"
unfolding \<epsilon>_def
using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.\<eta>\<epsilon>.\<epsilon>.is_natural_2
- hom\<^sub>C.comp_char G_def Faa'.F\<^sub>1_def G\<^sub>0_props hom\<^sub>C.dom_char hom\<^sub>C.cod_char
+ hom\<^sub>C.comp_char G_def Faa'.F\<^sub>1_def G\<^sub>0_props hom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C
apply simp
- by (metis (no_types, lifting) C.in_hhomI Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_reflects_arr hom\<^sub>C.arr_char
+ by (metis (no_types, lifting) C.in_hhomI Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_reflects_arr hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C
hom\<^sub>C.not_arr_null hom\<^sub>C.null_char)
show 2: "\<mu> \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) = \<epsilon> \<mu>"
unfolding \<epsilon>_def
using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.\<eta>\<epsilon>.\<epsilon>.is_natural_1
- hom\<^sub>C.comp_char G_def Faa'.F\<^sub>1_def G\<^sub>0_props hom\<^sub>C.dom_char hom\<^sub>C.cod_char
+ hom\<^sub>C.comp_char G_def Faa'.F\<^sub>1_def G\<^sub>0_props hom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C
apply simp
by (metis (no_types, lifting) C.in_hhomI Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_reflects_arr
- hom\<^sub>C.arr_char hom\<^sub>C.not_arr_null hom\<^sub>C.null_char)
+ hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.not_arr_null hom\<^sub>C.null_char)
show "G (F \<mu>) \<cdot>\<^sub>C C.inv (\<epsilon> (C.dom \<mu>)) = C.inv (\<epsilon> (C.cod \<mu>)) \<cdot>\<^sub>C \<mu>"
using assms 1 2 C.invert_opposite_sides_of_square by simp
qed
interpretation \<epsilon>: natural_isomorphism V\<^sub>C V\<^sub>C GF.map C.map \<epsilon>
using \<epsilon>_simps \<epsilon>_naturality \<epsilon>_def
by unfold_locales auto
interpretation GFG: composite_functor V\<^sub>D V\<^sub>C V\<^sub>C G \<open>GF.map\<close> ..
interpretation FGF: composite_functor V\<^sub>C V\<^sub>D V\<^sub>D F \<open>FG.map\<close> ..
interpretation Go\<eta>: natural_transformation V\<^sub>D V\<^sub>C G GFG.map \<open>G \<circ> \<eta>\<close>
proof -
have "G \<circ> D.map = G"
using G.as_nat_trans.natural_transformation_axioms by auto
moreover have "G \<circ> FG.map = GFG.map"
by auto
ultimately show "natural_transformation V\<^sub>D V\<^sub>C G GFG.map (G \<circ> \<eta>)"
using \<eta>.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms
horizontal_composite [of V\<^sub>D V\<^sub>D D.map FG.map \<eta> V\<^sub>C G G G] by simp
qed
interpretation \<eta>oF: natural_transformation V\<^sub>C V\<^sub>D F FGF.map \<open>\<eta> \<circ> F\<close>
using \<eta>.natural_transformation_axioms as_nat_trans.natural_transformation_axioms
horizontal_composite [of V\<^sub>C V\<^sub>D F F F V\<^sub>D D.map FG.map \<eta>]
by simp
interpretation \<epsilon>oG: natural_transformation V\<^sub>D V\<^sub>C GFG.map G \<open>\<epsilon> \<circ> G\<close>
using \<epsilon>.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms
horizontal_composite [of V\<^sub>D V\<^sub>C G G G V\<^sub>C GF.map C.map \<epsilon>]
by simp
interpretation Fo\<epsilon>: natural_transformation V\<^sub>C V\<^sub>D FGF.map F \<open>F \<circ> \<epsilon>\<close>
proof -
have "F \<circ> C.map = F"
using as_nat_trans.natural_transformation_axioms by auto
moreover have "F \<circ> GF.map = FGF.map"
by auto
ultimately show "natural_transformation V\<^sub>C V\<^sub>D FGF.map F (F \<circ> \<epsilon>)"
using \<epsilon>.natural_transformation_axioms as_nat_trans.natural_transformation_axioms
horizontal_composite [of V\<^sub>C V\<^sub>C GF.map C.map \<epsilon> V\<^sub>D F F F]
by simp
qed
interpretation \<epsilon>oG_Go\<eta>: vertical_composite V\<^sub>D V\<^sub>C G GFG.map G \<open>G \<circ> \<eta>\<close> \<open>\<epsilon> \<circ> G\<close> ..
interpretation Fo\<epsilon>_\<eta>oF: vertical_composite V\<^sub>C V\<^sub>D F FGF.map F \<open>\<eta> \<circ> F\<close> \<open>F \<circ> \<epsilon>\<close> ..
text \<open>
Bijectivity results in an ordinary adjunction between the vertical categories.
\<close>
lemma adjunction_\<eta>\<epsilon>:
shows "unit_counit_adjunction V\<^sub>C V\<^sub>D G F \<eta> \<epsilon>"
proof
show "\<epsilon>oG_Go\<eta>.map = G"
proof
fix \<nu>
have "\<not> D.arr \<nu> \<Longrightarrow> \<epsilon>oG_Go\<eta>.map \<nu> = G \<nu>"
unfolding \<epsilon>oG_Go\<eta>.map_def
by (simp add: G.is_extensional)
moreover have "D.arr \<nu> \<Longrightarrow> \<epsilon>oG_Go\<eta>.map \<nu> = G \<nu>"
proof -
assume \<nu>: "D.arr \<nu>"
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>'. \<guillemotleft>\<mu>' : G\<^sub>0 (src\<^sub>D \<nu>) \<rightarrow>\<^sub>C G\<^sub>0 (trg\<^sub>D \<nu>)\<guillemotright>\<close>
using \<nu> C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>'. \<guillemotleft>\<nu>' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>)) \<rightarrow>\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>))\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>G\<^sub>0 (src\<^sub>D \<nu>)\<close> \<open>G\<^sub>0 (trg\<^sub>D \<nu>)\<close>
using \<nu> G\<^sub>0_props by unfold_locales simp_all
show "\<epsilon>oG_Go\<eta>.map \<nu> = G \<nu>"
proof -
have 1: "C.arr (Faa'.G\<^sub>1 (D.cod \<nu>))"
- using \<nu> hom\<^sub>D.arr_char G_def
- by (metis Faa'.\<eta>\<epsilon>.F.preserves_reflects_arr G_in_hom(1) hom\<^sub>C.arr_char
+ using \<nu> hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C G_def
+ by (metis Faa'.\<eta>\<epsilon>.F.preserves_reflects_arr G_in_hom(1) hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C
hom\<^sub>C.inclusion hom\<^sub>D.cod_closed)
have 2: "D.arr (Faa'.\<eta> \<nu>)"
- using \<nu> hom\<^sub>D.arr_char
+ using \<nu> hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis D.in_hhomI D.obj_src D.obj_trg Faa'.\<eta>\<epsilon>.\<eta>.preserves_reflects_arr
G\<^sub>0_props(1) hom\<^sub>D.inclusion)
have 3: "src\<^sub>C (Faa'.G\<^sub>1 (D.cod \<nu>)) = G\<^sub>0 (src\<^sub>D \<nu>)"
using \<nu>
by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(2))
have 4: "trg\<^sub>C (Faa'.G\<^sub>1 (D.cod \<nu>)) = G\<^sub>0 (trg\<^sub>D \<nu>)"
using \<nu>
by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(3))
have 5: "hom\<^sub>D.arr \<nu>"
- using \<nu> hom\<^sub>D.arr_char G\<^sub>0_props by simp
+ using \<nu> hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C G\<^sub>0_props by simp
have "\<epsilon>oG_Go\<eta>.map \<nu> = \<epsilon> (G (D.cod \<nu>)) \<cdot>\<^sub>C G (\<eta> \<nu>)"
using \<nu> \<epsilon>oG_Go\<eta>.map_def by simp
also have "... = Faa'.\<epsilon> (Faa'.G\<^sub>1 (D.cod \<nu>)) \<cdot>\<^sub>C Faa'.G\<^sub>1 (Faa'.\<eta> \<nu>)"
using \<nu> 1 2 3 4 \<eta>_def \<epsilon>_def G_def G_simps(2-3) \<eta>_simps(2-3) by auto
also have "... = hom\<^sub>C.comp ((Faa'.\<epsilon> \<circ> Faa'.G\<^sub>1) (hom\<^sub>D.cod \<nu>)) ((Faa'.G\<^sub>1 \<circ> Faa'.\<eta>) \<nu>)"
proof -
have "Faa'.\<epsilon> (Faa'.G\<^sub>1 (D.cod \<nu>)) \<cdot>\<^sub>C Faa'.G\<^sub>1 (Faa'.\<eta> \<nu>) = Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map \<nu>"
proof -
have "C.seq (Faa'.\<epsilon> (Faa'.G\<^sub>1 (hom\<^sub>D.cod \<nu>))) (Faa'.G\<^sub>1 (Faa'.\<eta> \<nu>))"
proof -
have "hom\<^sub>C.seq (Faa'.\<epsilon> (Faa'.G\<^sub>1 (hom\<^sub>D.cod \<nu>))) (Faa'.G\<^sub>1 (Faa'.\<eta> \<nu>))"
- using \<nu> 5 hom\<^sub>C.arr_char hom\<^sub>C.seq_char Faa'.\<eta>\<epsilon>.F.preserves_arr
+ using \<nu> 5 hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.seq_char\<^sub>S\<^sub>b\<^sub>C Faa'.\<eta>\<epsilon>.F.preserves_arr
apply (intro hom\<^sub>C.seqI)
apply auto
using Faa'.\<eta>\<epsilon>.\<epsilon>.preserves_reflects_arr hom\<^sub>C.inclusion hom\<^sub>D.arr_cod
by presburger
thus ?thesis
- using hom\<^sub>C.seq_char by simp
+ using hom\<^sub>C.seq_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
moreover have "D.in_hhom (Faa'.\<eta> \<nu>) (F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>))) (F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>)))"
using \<nu> 5 Faa'.\<eta>\<epsilon>.\<eta>.preserves_reflects_arr hom\<^sub>D.arrE by blast
moreover have "D.in_hhom (D.cod \<nu>) (F\<^sub>0 (G\<^sub>0 (src\<^sub>D \<nu>))) (F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \<nu>)))"
using \<nu> 5 hom\<^sub>D.arrE hom\<^sub>D.cod_closed by blast
ultimately show ?thesis
- using \<nu> 5 Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map_def hom\<^sub>D.arr_char hom\<^sub>C.comp_char hom\<^sub>D.cod_char
+ using \<nu> 5 Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map_def hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.comp_char hom\<^sub>D.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
qed
thus ?thesis
using \<nu> 5 Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map_def by simp
qed
also have "... = G \<nu>"
using \<nu> 5 G_def Faa'.\<eta>\<epsilon>.triangle_F Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map_simp_1 [of \<nu>] by simp
finally show ?thesis by simp
qed
qed
ultimately show "\<epsilon>oG_Go\<eta>.map \<nu> = G \<nu>" by auto
qed
show "Fo\<epsilon>_\<eta>oF.map = F"
proof
fix \<mu>
have "\<not> C.arr \<mu> \<Longrightarrow> Fo\<epsilon>_\<eta>oF.map \<mu> = F \<mu>"
unfolding Fo\<epsilon>_\<eta>oF.map_def
by (simp add: is_extensional)
moreover have "C.arr \<mu> \<Longrightarrow> Fo\<epsilon>_\<eta>oF.map \<mu> = F \<mu>"
proof -
assume \<mu>: "C.arr \<mu>"
interpret hom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>'. \<guillemotleft>\<mu>' : src\<^sub>C \<mu> \<rightarrow>\<^sub>C trg\<^sub>C \<mu>\<guillemotright>\<close>
using \<mu> C.hhom_is_subcategory by simp
interpret hom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<nu>. \<guillemotleft>\<nu> : F\<^sub>0 (src\<^sub>C \<mu>) \<rightarrow>\<^sub>D F\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>\<close>
using \<mu> D.in_hhom_def D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi> \<open>src\<^sub>C \<mu>\<close> \<open>trg\<^sub>C \<mu>\<close>
using \<mu> by unfold_locales auto
show "Fo\<epsilon>_\<eta>oF.map \<mu> = F \<mu>"
proof -
have 1: "hom\<^sub>C.arr \<mu>"
- using \<mu> hom\<^sub>C.arr_char by simp
+ using \<mu> hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have "Fo\<epsilon>_\<eta>oF.map \<mu> = F (\<epsilon> (C.cod \<mu>)) \<cdot>\<^sub>D \<eta> (F \<mu>)"
using \<mu> Fo\<epsilon>_\<eta>oF.map_def by simp
also have "... = Faa'.F\<^sub>1 (Faa'.\<epsilon> (C.cod \<mu>)) \<cdot>\<^sub>D Faa'.\<eta> (Faa'.F\<^sub>1 \<mu>)"
unfolding \<eta>_def \<epsilon>_def G_def
using \<mu> G\<^sub>0_props Faa'.F\<^sub>1_def
by auto
also have "... = hom\<^sub>D.comp ((Faa'.F\<^sub>1 \<circ> Faa'.\<epsilon>) (hom\<^sub>C.cod \<mu>)) ((Faa'.\<eta> \<circ> Faa'.F\<^sub>1) \<mu>)"
proof -
have 2: "hom\<^sub>C.arr (C.cod \<mu>)"
- using 1 hom\<^sub>C.cod_char [of \<mu>] hom\<^sub>C.arr_cod hom\<^sub>C.dom_char hom\<^sub>C.cod_char
+ using 1 hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C [of \<mu>] hom\<^sub>C.arr_cod hom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "D.seq (Faa'.F\<^sub>1 (Faa'.\<epsilon> (C.cod \<mu>))) (Faa'.\<eta> (Faa'.F\<^sub>1 \<mu>))"
proof -
have "hom\<^sub>D.seq (Faa'.F\<^sub>1 (Faa'.\<epsilon> (C.cod \<mu>))) (Faa'.\<eta> (Faa'.F\<^sub>1 \<mu>))"
- using \<mu> 1 2 hom\<^sub>D.arr_char hom\<^sub>D.seq_char Faa'.\<eta>\<epsilon>.G.preserves_arr
+ using \<mu> 1 2 hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>D.seq_char\<^sub>S\<^sub>b\<^sub>C Faa'.\<eta>\<epsilon>.G.preserves_arr
hom\<^sub>C.cod_simp hom\<^sub>C.dom_cod
apply (intro hom\<^sub>D.seqI)
by auto metis
thus ?thesis
- using hom\<^sub>D.seq_char by simp
+ using hom\<^sub>D.seq_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
ultimately show ?thesis
- using 1 hom\<^sub>D.comp_char hom\<^sub>C.dom_char hom\<^sub>C.cod_char by simp
+ using 1 hom\<^sub>D.comp_char hom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
also have "... = Faa'.\<eta>\<epsilon>.G\<epsilon>o\<eta>G.map \<mu>"
unfolding Faa'.\<eta>\<epsilon>.G\<epsilon>o\<eta>G.map_def
using 1 by simp
also have "... = F \<mu>"
- using \<mu> Faa'.\<eta>\<epsilon>.triangle_G hom\<^sub>C.arr_char Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map_def
- hom\<^sub>D.arr_char hom\<^sub>C.comp_char hom\<^sub>D.comp_char Faa'.F\<^sub>1_def
+ using \<mu> Faa'.\<eta>\<epsilon>.triangle_G hom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C Faa'.\<eta>\<epsilon>.\<epsilon>FoF\<eta>.map_def
+ hom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hom\<^sub>C.comp_char hom\<^sub>D.comp_char Faa'.F\<^sub>1_def
by simp
finally show ?thesis by simp
qed
qed
ultimately show "Fo\<epsilon>_\<eta>oF.map \<mu> = F \<mu>" by auto
qed
qed
interpretation \<eta>\<epsilon>: unit_counit_adjunction V\<^sub>C V\<^sub>D G F \<eta> \<epsilon>
using adjunction_\<eta>\<epsilon> by simp
text \<open>
We now use the adjunction between the vertical categories to define the
compositors for \<open>G\<close>. Without the bijectivity assumption, we would only obtain
\<open>\<eta>\<close> and \<open>\<epsilon>\<close> as pseudonatural equivalences, rather than natural isomorphisms,
which would make everything more complicated.
\<close>
definition \<Phi>\<^sub>G\<^sub>0
where "\<Phi>\<^sub>G\<^sub>0 f f' = C.inv (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
lemma \<Phi>\<^sub>G\<^sub>0_in_hom:
assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'"
shows "\<guillemotleft>\<Phi>\<^sub>G\<^sub>0 f f' : G f \<star>\<^sub>C G f' \<Rightarrow>\<^sub>C G (f \<star>\<^sub>D f')\<guillemotright>"
proof (unfold \<Phi>\<^sub>G\<^sub>0_def, intro C.inv_in_hom)
show 1: "\<guillemotleft>\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')) :
G (f \<star>\<^sub>D f') \<Rightarrow>\<^sub>C G f \<star>\<^sub>C G f'\<guillemotright>"
proof
show "\<guillemotleft>\<epsilon> (G f \<star>\<^sub>C G f') : G (F (G f \<star>\<^sub>C G f')) \<Rightarrow>\<^sub>C G f \<star>\<^sub>C G f'\<guillemotright>"
using assms \<epsilon>_in_hom by auto
show "\<guillemotleft>G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')) : G (f \<star>\<^sub>D f') \<Rightarrow>\<^sub>C G (F (G f \<star>\<^sub>C G f'))\<guillemotright>"
proof -
have "\<guillemotleft>\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f') : f \<star>\<^sub>D f' \<Rightarrow>\<^sub>D F (G f \<star>\<^sub>C G f')\<guillemotright>"
using assms \<eta>_in_hom
by (intro D.comp_in_homI') auto
thus ?thesis by auto
qed
qed
show "C.iso (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
proof (intro C.isos_compose)
show "C.iso (G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
proof -
have "D.iso (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f'))"
using assms \<eta>_in_hom
by (intro D.isos_compose D.seqI) auto
thus ?thesis by simp
qed
show "C.iso (\<epsilon> (G f \<star>\<^sub>C G f'))"
using assms \<epsilon>_in_hom by simp
show "C.seq (\<epsilon> (G f \<star>\<^sub>C G f')) (G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
using 1 by auto
qed
qed
lemma iso_\<Phi>\<^sub>G\<^sub>0:
assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'"
shows "C.iso (\<Phi>\<^sub>G\<^sub>0 f f')"
proof (unfold \<Phi>\<^sub>G\<^sub>0_def, intro C.iso_inv_iso C.isos_compose)
show 1: "C.iso (G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
using assms \<eta>_in_hom D.in_hhom_def cmp_simps'(1) cmp_simps(4) by auto
show "C.iso (\<epsilon> (G f \<star>\<^sub>C G f'))"
using assms \<epsilon>_in_hom by simp
show "C.seq (\<epsilon> (G f \<star>\<^sub>C G f')) (G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
using assms 1 by force
qed
lemma \<Phi>\<^sub>G\<^sub>0_naturality:
assumes "D.arr \<nu>" and "D.arr \<nu>'" and "src\<^sub>D \<nu> = trg\<^sub>D \<nu>'"
shows "\<Phi>\<^sub>G\<^sub>0 (D.cod \<nu>) (D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>') =
G (\<nu> \<star>\<^sub>D \<nu>') \<cdot>\<^sub>C \<Phi>\<^sub>G\<^sub>0 (D.dom \<nu>) (D.dom \<nu>')"
proof -
let ?X = "\<epsilon> (G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>')) \<cdot>\<^sub>C
G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')) \<cdot>\<^sub>D (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>')))"
have iso_X: "C.iso ?X"
- using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
by (intro C.isos_compose) auto
have "\<Phi>\<^sub>G\<^sub>0 (D.cod \<nu>) (D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>') = C.inv ?X \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>')"
unfolding \<Phi>\<^sub>G\<^sub>0_def by simp
also have "... = (C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')) \<cdot>\<^sub>D
(\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>')))) \<cdot>\<^sub>C
(G \<nu> \<star>\<^sub>C G \<nu>')"
using assms iso_X \<eta>_in_hom \<epsilon>_in_hom
by (simp add: C.inv_comp_left(1))
also have "... = (C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>'))) \<cdot>\<^sub>C
G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>')))) \<cdot>\<^sub>C
(G \<nu> \<star>\<^sub>C G \<nu>')"
proof -
have "G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')) \<cdot>\<^sub>D (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>'))) =
G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>'))) \<cdot>\<^sub>C G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>'))"
using assms iso_X \<eta>_in_hom by fast
thus ?thesis by simp
qed
also have "... = C.inv (G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>'))) \<cdot>\<^sub>C
(G \<nu> \<star>\<^sub>C G \<nu>')"
proof -
have "C.iso (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>'))) \<cdot>\<^sub>C G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>')))"
- using assms iso_X \<eta>_in_hom C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ using assms iso_X \<eta>_in_hom C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
by (intro C.isos_compose C.seqI) auto
hence "C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>'))) \<cdot>\<^sub>C
G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>')))
= C.inv (G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>'))))"
using assms \<eta>_in_hom
by (simp add: C.inv_comp_right(1))
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = G (D.inv (\<eta> (D.cod \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>'))) \<cdot>\<^sub>C
(G \<nu> \<star>\<^sub>C G \<nu>')"
proof -
have "C.inv (G (\<eta> (D.cod \<nu>) \<star>\<^sub>D \<eta> (D.cod \<nu>'))) =
G (D.inv (\<eta> (D.cod \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>')))"
using assms \<eta>_in_hom D.ide_cod D.iso_hcomp D.src_cod D.trg_cod G.preserves_inv
\<eta>_simps(2-3) D.inv_hcomp
by (metis D.arr_cod)
thus ?thesis by simp
qed
also have "... = G (D.inv (\<eta> (D.cod \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C
(C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')))) \<cdot>\<^sub>C
G (F (G \<nu> \<star>\<^sub>C G \<nu>'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')))"
proof -
have "C.inv (\<epsilon> (G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>'))) \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>')
= G (F (G \<nu> \<star>\<^sub>C G \<nu>')) \<cdot>\<^sub>C C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')))"
proof -
have "G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>') = C.dom (G \<nu> \<star>\<^sub>C G \<nu>')"
by (simp add: assms)
thus ?thesis
using assms C.hcomp_simps(4) C.hseqI' G.preserves_cod G.preserves_hseq
\<epsilon>_naturality(3)
by presburger
qed
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = (G (D.inv (\<eta> (D.cod \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C
G (F (G \<nu>) \<star>\<^sub>D F (G \<nu>'))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')))"
proof -
have "C.inv (G (\<Phi> (G (D.cod \<nu>), G (D.cod \<nu>')))) \<cdot>\<^sub>C G (F (G \<nu> \<star>\<^sub>C G \<nu>')) =
G (D.inv (\<Phi> (C.cod (G \<nu>), C.cod (G \<nu>'))) \<cdot>\<^sub>D F (G \<nu> \<star>\<^sub>C G \<nu>'))"
using assms by (simp add: G.preserves_inv)
also have "... = G ((F (G \<nu>) \<star>\<^sub>D F (G \<nu>')) \<cdot>\<^sub>D D.inv (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>'))))"
- using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
\<Phi>.inv_naturality [of "(G \<nu>, G \<nu>')"]
by simp
also have "... = G (F (G \<nu>) \<star>\<^sub>D F (G \<nu>')) \<cdot>\<^sub>C
G (D.inv (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>'))))"
using assms by simp
finally show ?thesis
using C.comp_assoc by simp
qed
also have "... = G (\<nu> \<star>\<^sub>D \<nu>') \<cdot>\<^sub>C
G (D.inv (\<eta> (D.dom \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.dom \<nu>'))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')))"
proof -
have "G (D.inv (\<eta> (D.cod \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>'))) \<cdot>\<^sub>C G (F (G \<nu>) \<star>\<^sub>D F (G \<nu>')) =
G ((D.inv (\<eta> (D.cod \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>'))) \<cdot>\<^sub>D (F (G \<nu>) \<star>\<^sub>D F (G \<nu>')))"
using assms by simp
also have "... = G (D.inv (\<eta> (D.cod \<nu>)) \<cdot>\<^sub>D F (G \<nu>)
\<star>\<^sub>D D.inv (\<eta> (D.cod \<nu>')) \<cdot>\<^sub>D F (G \<nu>'))"
using assms D.interchange by simp
also have "... = G (\<nu> \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>)) \<star>\<^sub>D \<nu>' \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>')))"
using assms \<eta>_naturality by simp
also have "... = G ((\<nu> \<star>\<^sub>D \<nu>') \<cdot>\<^sub>D (D.inv (\<eta> (D.dom \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.dom \<nu>'))))"
using assms D.interchange by simp
also have "... = G (\<nu> \<star>\<^sub>D \<nu>') \<cdot>\<^sub>C G (D.inv (\<eta> (D.dom \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.dom \<nu>')))"
using assms by simp
finally show ?thesis
using C.comp_assoc by simp
qed
also have "... = G (\<nu> \<star>\<^sub>D \<nu>') \<cdot>\<^sub>C \<Phi>\<^sub>G\<^sub>0 (D.dom \<nu>) (D.dom \<nu>')"
proof -
have "G (D.inv (\<eta> (D.dom \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.dom \<nu>'))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')))
= C.inv (G (\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')))"
proof -
have "G (D.inv (\<eta> (D.dom \<nu>)) \<star>\<^sub>D D.inv (\<eta> (D.dom \<nu>'))) =
C.inv (G (\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>')))"
using assms cmp_components_are_iso
by (metis D.arr_dom D.ide_dom D.inv_hcomp D.iso_hcomp D.src_dom D.trg_dom
G.preserves_inv \<eta>_simps(2) \<eta>_simps(3) \<eta>_simps(6))
moreover have "G (D.inv (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')))) =
C.inv (G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>'))))"
using assms cmp_components_are_iso
by (simp add: G.preserves_inv)
ultimately show ?thesis by simp
qed
also have "... = C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')) \<cdot>\<^sub>C
G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>'))) \<cdot>\<^sub>C
G (\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>')))"
proof -
have "C.iso (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')) \<cdot>\<^sub>C
G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>'))) \<cdot>\<^sub>C
G (\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>')))"
- using assms \<eta>_in_hom \<epsilon>_in_hom cmp_simps'(1,4) C.VV.cod_char
+ using assms \<eta>_in_hom \<epsilon>_in_hom cmp_simps'(1,4) C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by (intro C.isos_compose C.seqI) auto
thus ?thesis
using assms cmp_components_are_iso C.comp_assoc C.inv_comp_left by simp
qed
also have "... = C.inv (\<epsilon> (G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')) \<cdot>\<^sub>C
G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')) \<cdot>\<^sub>D
(\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>'))))"
proof -
have "G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>'))) \<cdot>\<^sub>C G (\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>')) =
G (\<Phi> (C.dom (G \<nu>), C.dom (G \<nu>')) \<cdot>\<^sub>D (\<eta> (D.dom \<nu>) \<star>\<^sub>D \<eta> (D.dom \<nu>')))"
using assms cmp_simps'(1,4) by auto
thus ?thesis by simp
qed
also have "... = \<Phi>\<^sub>G\<^sub>0 (D.dom \<nu>) (D.dom \<nu>')"
unfolding \<Phi>\<^sub>G\<^sub>0_def
using assms by simp
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
interpretation \<Phi>\<^sub>G: transformation_by_components D.VV.comp V\<^sub>C H\<^sub>DoGG.map GoH\<^sub>D.map
\<open>\<lambda>fg. \<Phi>\<^sub>G\<^sub>0 (fst fg) (snd fg)\<close>
- using D.VV.ide_char D.VV.arr_char \<Phi>\<^sub>G\<^sub>0_in_hom \<Phi>\<^sub>G\<^sub>0_naturality G.FF_def
- D.VV.dom_char D.VV.cod_char
+ using D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>G\<^sub>0_in_hom \<Phi>\<^sub>G\<^sub>0_naturality G.FF_def
+ D.VV.dom_char\<^sub>S\<^sub>b\<^sub>C D.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by unfold_locales auto
interpretation \<Phi>\<^sub>G: natural_isomorphism D.VV.comp V\<^sub>C H\<^sub>DoGG.map GoH\<^sub>D.map \<Phi>\<^sub>G.map
- using \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char iso_\<Phi>\<^sub>G\<^sub>0 by unfold_locales simp
+ using \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C iso_\<Phi>\<^sub>G\<^sub>0 by unfold_locales simp
abbreviation \<Phi>\<^sub>G
where "\<Phi>\<^sub>G \<equiv> \<Phi>\<^sub>G.map"
lemma \<Phi>\<^sub>G_in_hom [intro]:
assumes "D.arr \<nu>" and "D.arr \<nu>'" and "src\<^sub>D \<nu> = trg\<^sub>D \<nu>'"
shows "C.in_hhom (\<Phi>\<^sub>G.map (\<nu>, \<nu>')) (src\<^sub>C (G (D.dom \<nu>'))) (trg\<^sub>C (G (D.cod \<nu>)))"
and "\<guillemotleft>\<Phi>\<^sub>G.map (\<nu>, \<nu>') : G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>') \<Rightarrow>\<^sub>C G (D.cod \<nu> \<star>\<^sub>D D.cod \<nu>')\<guillemotright>"
proof -
show "\<guillemotleft>\<Phi>\<^sub>G.map (\<nu>, \<nu>') : G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>') \<Rightarrow>\<^sub>C G (D.cod \<nu> \<star>\<^sub>D D.cod \<nu>')\<guillemotright>"
proof -
have "\<Phi>\<^sub>G.map (\<nu>, \<nu>') = \<Phi>\<^sub>G\<^sub>0 (D.cod \<nu>) (D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>')"
- using assms D.VV.arr_char \<Phi>\<^sub>G.map_def \<Phi>\<^sub>G\<^sub>0_in_hom G.FF_def D.VV.cod_char
+ using assms D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>G.map_def \<Phi>\<^sub>G\<^sub>0_in_hom G.FF_def D.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "\<guillemotleft>\<Phi>\<^sub>G\<^sub>0 (D.cod \<nu>) (D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>') :
G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>') \<Rightarrow>\<^sub>C G (D.cod \<nu> \<star>\<^sub>D D.cod \<nu>')\<guillemotright>"
using assms \<Phi>\<^sub>G\<^sub>0_in_hom by (intro C.comp_in_homI) auto
ultimately show ?thesis by simp
qed
thus "C.in_hhom (\<Phi>\<^sub>G.map (\<nu>, \<nu>')) (src\<^sub>C (G (D.dom \<nu>'))) (trg\<^sub>C (G (D.cod \<nu>)))"
using assms C.vconn_implies_hpar(1-2) by auto
qed
lemma \<Phi>\<^sub>G_simps [simp]:
assumes "D.arr \<nu>" and "D.arr \<nu>'" and "src\<^sub>D \<nu> = trg\<^sub>D \<nu>'"
shows "C.arr (\<Phi>\<^sub>G.map (\<nu>, \<nu>'))"
and "src\<^sub>C (\<Phi>\<^sub>G.map (\<nu>, \<nu>')) = src\<^sub>C (G (D.dom \<nu>'))"
and "trg\<^sub>C (\<Phi>\<^sub>G.map (\<nu>, \<nu>')) = trg\<^sub>C (G (D.cod \<nu>))"
and "C.dom (\<Phi>\<^sub>G.map (\<nu>, \<nu>')) = G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>')"
and "C.cod (\<Phi>\<^sub>G.map (\<nu>, \<nu>')) = G (D.cod \<nu> \<star>\<^sub>D D.cod \<nu>')"
using assms \<Phi>\<^sub>G_in_hom
apply auto
by fast+
lemma \<Phi>\<^sub>G_map_simp_ide:
assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'"
shows "\<Phi>\<^sub>G.map (f, f') = G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> f')) \<cdot>\<^sub>C G (D.inv (\<Phi> (G f, G f'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))"
proof -
have "\<Phi>\<^sub>G.map (f, f') = C.inv (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))"
- using assms \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char \<Phi>\<^sub>G\<^sub>0_def G.FF_def by auto
+ using assms \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>G\<^sub>0_def G.FF_def by auto
also have "... = C.inv (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f')) \<cdot>\<^sub>C G (\<eta> f \<star>\<^sub>D \<eta> f'))"
- using assms D.VV.ide_char D.VV.arr_char cmp_simps(1,4) by simp
+ using assms D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_simps(1,4) by simp
also have "... = C.inv (G (\<eta> f \<star>\<^sub>D \<eta> f')) \<cdot>\<^sub>C C.inv (G (\<Phi> (G f, G f'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))"
proof -
have "C.iso (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f')) \<cdot>\<^sub>C G (\<eta> f \<star>\<^sub>D \<eta> f'))"
- using assms C.VV.ide_char C.VV.arr_char FF_def
+ using assms C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
by (intro C.isos_compose C.seqI) auto
thus ?thesis
using assms C.inv_comp_left(1-2) cmp_components_are_iso C.comp_assoc by simp
qed
also have "... = G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> f')) \<cdot>\<^sub>C G (D.inv (\<Phi> (G f, G f'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))"
using assms cmp_components_are_iso D.ideD(1) D.inv_hcomp D.iso_hcomp
G.preserves_ide G.preserves_inv G_simps(2-3) \<eta>_simps(2-3,6)
by metis
finally show ?thesis by blast
qed
lemma \<eta>_hcomp:
assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'"
shows "\<eta> (f \<star>\<^sub>D f') = F (\<Phi>\<^sub>G.map (f, f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')"
proof -
have 1: "\<guillemotleft>F (\<Phi>\<^sub>G.map (f, f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f') :
f \<star>\<^sub>D f' \<Rightarrow>\<^sub>D F (G (f \<star>\<^sub>D f'))\<guillemotright>"
using assms by fastforce
have 2: "\<guillemotleft>\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f')) \<cdot>\<^sub>C G (\<eta> f \<star>\<^sub>D \<eta> f') :
G (f \<star>\<^sub>D f') \<Rightarrow>\<^sub>C G f \<star>\<^sub>C G f'\<guillemotright>"
using assms G.preserves_hom cmp_in_hom(2)
by (intro C.comp_in_homI) auto
have "F (\<Phi>\<^sub>G.map (f, f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')
= F (C.inv (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')))) \<cdot>\<^sub>D
\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')"
- using assms \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char \<Phi>\<^sub>G\<^sub>0_def by simp
+ using assms \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>G\<^sub>0_def by simp
also have "... = F (C.inv (G (\<eta> f \<star>\<^sub>D \<eta> f')) \<cdot>\<^sub>C C.inv (G (\<Phi> (G f, G f'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))) \<cdot>\<^sub>D
\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')"
proof -
have "C.inv (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f'))) =
C.inv (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f')) \<cdot>\<^sub>C G (\<eta> f \<star>\<^sub>D \<eta> f'))"
using assms 1 by force
also have "... = C.inv (G (\<eta> f \<star>\<^sub>D \<eta> f')) \<cdot>\<^sub>C C.inv (G (\<Phi> (G f, G f'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))"
proof -
have "C.iso (\<epsilon> (G f \<star>\<^sub>C G f') \<cdot>\<^sub>C G (\<Phi> (G f, G f')) \<cdot>\<^sub>C G (\<eta> f \<star>\<^sub>D \<eta> f'))"
using assms 2 by (intro C.isos_compose) auto
thus ?thesis
using assms 1 C.inv_comp_left C.comp_assoc by simp
qed
finally show ?thesis by simp
qed
also have "... = F (C.inv (G (\<eta> f \<star>\<^sub>D \<eta> f'))) \<cdot>\<^sub>D
F (C.inv (G (\<Phi> (G f, G f')))) \<cdot>\<^sub>D F (C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))) \<cdot>\<^sub>D
\<Phi> (G f, G f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')"
proof -
have "C.arr ((C.inv (G (\<eta> f \<star>\<^sub>D \<eta> f')) \<cdot>\<^sub>C C.inv (G (\<Phi> (G f, G f'))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G f'))))"
- using assms 1 2 cmp_components_are_iso C.VV.ide_char C.VV.arr_char FF_def
+ using assms 1 2 cmp_components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
by (intro C.seqI) auto
thus ?thesis
using C.inv_comp_left D.comp_assoc by auto
qed
also have "... = D.inv (F (G (\<eta> f \<star>\<^sub>D \<eta> f'))) \<cdot>\<^sub>D
D.inv (F (G (\<Phi> (G f, G f')))) \<cdot>\<^sub>D
D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f'))) \<cdot>\<^sub>D
\<Phi> (G f, G f') \<cdot>\<^sub>D
(\<eta> f \<star>\<^sub>D \<eta> f')"
using assms by (simp add: preserves_inv)
also have "... = D.inv ((\<eta> (F (G f) \<star>\<^sub>D F (G f'))) \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D
D.inv (\<eta> (f \<star>\<^sub>D f'))) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f \<star>\<^sub>C G f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D
D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f')))) \<cdot>\<^sub>D
D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f'))) \<cdot>\<^sub>D
\<Phi> (G f, G f') \<cdot>\<^sub>D
(\<eta> f \<star>\<^sub>D \<eta> f')"
proof -
have *: "\<And>\<nu>. D.arr \<nu> \<Longrightarrow> \<eta> (D.cod \<nu>) \<cdot>\<^sub>D \<nu> \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>)) = F (G \<nu>)"
by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
\<eta>_naturality(1-2) \<eta>_simps(1,6))
have "F (G (\<eta> f \<star>\<^sub>D \<eta> f')) =
\<eta> (F (G f) \<star>\<^sub>D F (G f')) \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D D.inv (\<eta> (f \<star>\<^sub>D f'))"
using assms * [of "\<eta> f \<star>\<^sub>D \<eta> f'"] by simp
moreover have "F (G (\<Phi> (G f, G f'))) =
\<eta> (F (G f \<star>\<^sub>C G f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f')))"
using assms 1 * [of "\<Phi> (G f, G f')"] cmp_simps(5) by fastforce
ultimately show ?thesis by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D f') \<cdot>\<^sub>D
D.inv (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D
((D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f')))) \<cdot>\<^sub>D (\<eta> (F (G f) \<star>\<^sub>D F (G f')))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G f')) \<cdot>\<^sub>D
((D.inv (\<eta> (F (G f \<star>\<^sub>C G f')))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f')))) \<cdot>\<^sub>D
\<Phi> (G f, G f')) \<cdot>\<^sub>D
(\<eta> f \<star>\<^sub>D \<eta> f')"
proof -
have "D.inv ((\<eta> (F (G f) \<star>\<^sub>D F (G f'))) \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D D.inv (\<eta> (f \<star>\<^sub>D f')))
= \<eta> (f \<star>\<^sub>D f') \<cdot>\<^sub>D D.inv (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f')))"
proof -
have "D.iso (\<eta> (F (G f) \<star>\<^sub>D F (G f')) \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D D.inv (\<eta> (f \<star>\<^sub>D f')))"
using assms by (intro D.isos_compose) auto
thus ?thesis
using assms D.inv_comp_left D.comp_assoc by simp
qed
moreover have
"D.inv (\<eta> (F (G f \<star>\<^sub>C G f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f')))) =
\<eta> (F (G f) \<star>\<^sub>D F (G f')) \<cdot>\<^sub>D D.inv (\<Phi> (G f, G f')) \<cdot>\<^sub>D D.inv (\<eta> (F (G f \<star>\<^sub>C G f')))"
proof -
have "D.iso (\<eta> (F (G f \<star>\<^sub>C G f')) \<cdot>\<^sub>D \<Phi> (G f, G f') \<cdot>\<^sub>D D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f'))))"
proof -
have 1: "D.seq (\<Phi> (G f, G f')) (D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f'))))"
using assms by (intro D.seqI) auto
moreover have "D.seq (\<eta> (F (G f \<star>\<^sub>C G f')))
(\<Phi> (G f, G f') \<cdot>\<^sub>D D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f'))))"
using assms 1 by auto
ultimately show ?thesis
using assms cmp_components_are_iso \<eta>_in_hom
by (intro D.isos_compose) auto
qed
thus ?thesis
using assms D.comp_inv_arr' D.comp_assoc D.inv_comp_left cmp_components_are_iso
by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D f') \<cdot>\<^sub>D
D.inv (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D
(D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f'))) \<cdot>\<^sub>D (\<eta> (F (G f) \<star>\<^sub>D F (G f')))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G f')) \<cdot>\<^sub>D \<Phi> (G f, G f')) \<cdot>\<^sub>D
(\<eta> f \<star>\<^sub>D \<eta> f')"
proof -
have "(D.inv (\<eta> (F (G f \<star>\<^sub>C G f')))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f'))) \<cdot>\<^sub>D \<Phi> (G f, G f')
= F (G f \<star>\<^sub>C G f') \<cdot>\<^sub>D \<Phi> (G f, G f')"
proof -
have "(D.inv (\<eta> (F (G f \<star>\<^sub>C G f')))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f')))
= D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f')) \<cdot>\<^sub>D \<eta> (F (G f \<star>\<^sub>C G f')))"
using assms by (simp add: D.inv_comp)
also have "... = F (G f \<star>\<^sub>C G f')"
using assms \<eta>\<epsilon>.triangle_G Fo\<epsilon>_\<eta>oF.map_simp_ide by fastforce
finally show ?thesis using D.comp_assoc by metis
qed
also have "... = \<Phi> (G f, G f')"
using assms D.comp_cod_arr cmp_in_hom(2) [of "G f" "G f'"] by auto
finally have "(D.inv (\<eta> (F (G f \<star>\<^sub>C G f')))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G f \<star>\<^sub>C G f'))) \<cdot>\<^sub>D
\<Phi> (G f, G f')
= \<Phi> (G f, G f')"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D f') \<cdot>\<^sub>D
D.inv (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D
(D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G f'))) \<cdot>\<^sub>D (\<eta> (F (G f) \<star>\<^sub>D F (G f')))) \<cdot>\<^sub>D
(\<eta> f \<star>\<^sub>D \<eta> f')"
using assms D.comp_cod_arr D.comp_inv_arr' cmp_components_are_iso by simp
also have "... = \<eta> (f \<star>\<^sub>D f') \<cdot>\<^sub>D D.inv (\<eta> f \<star>\<^sub>D \<eta> f') \<cdot>\<^sub>D (\<eta> f \<star>\<^sub>D \<eta> f')"
using assms \<eta>_in_hom D.comp_inv_arr' D.comp_cod_arr by simp
also have "... = \<eta> (f \<star>\<^sub>D f')"
using assms D.comp_inv_arr' [of "\<eta> f \<star>\<^sub>D \<eta> f'"] D.comp_arr_dom by simp
finally show ?thesis by simp
qed
lemma \<epsilon>_hcomp:
assumes "C.ide g" and "C.ide g'" and "src\<^sub>C g = trg\<^sub>C g'"
shows "\<epsilon> (g \<star>\<^sub>C g') = (\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>C C.inv (\<Phi>\<^sub>G.map (F g, F g')) \<cdot>\<^sub>C C.inv (G (\<Phi> (g, g')))"
proof -
have 1: "\<guillemotleft>\<epsilon> (G (F g) \<star>\<^sub>C G (F g')) \<cdot>\<^sub>C G (\<Phi> (G (F g), G (F g')) \<cdot>\<^sub>D (\<eta> (F g) \<star>\<^sub>D \<eta> (F g'))) :
G (F g \<star>\<^sub>D F g') \<Rightarrow>\<^sub>C G (F g) \<star>\<^sub>C G (F g')\<guillemotright>"
using assms by fastforce
have "(\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>C C.inv (\<Phi>\<^sub>G.map (F g, F g')) \<cdot>\<^sub>C C.inv (G (\<Phi> (g, g')))
= (\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>C
C.inv (C.inv (\<epsilon> (G (F g) \<star>\<^sub>C G (F g')) \<cdot>\<^sub>C G (\<Phi> (G (F g), G (F g')) \<cdot>\<^sub>D
(\<eta> (F g) \<star>\<^sub>D \<eta> (F g'))))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (g, g')))"
- using assms \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char \<Phi>\<^sub>G\<^sub>0_def by simp
+ using assms \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>\<^sub>G\<^sub>0_def by simp
also have "... = ((\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>C \<epsilon> (G (F g) \<star>\<^sub>C G (F g'))) \<cdot>\<^sub>C
G (\<Phi> (G (F g), G (F g')) \<cdot>\<^sub>D (\<eta> (F g) \<star>\<^sub>D \<eta> (F g'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (g, g')))"
proof -
have "C.iso (\<epsilon> (G (F g) \<star>\<^sub>C G (F g')) \<cdot>\<^sub>C G (\<Phi> (G (F g), G (F g')) \<cdot>\<^sub>D
(\<eta> (F g) \<star>\<^sub>D \<eta> (F g'))))"
using assms 1
by (intro C.isos_compose D.isos_compose G.preserves_iso) auto
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C
(G (F (\<epsilon> g \<star>\<^sub>C \<epsilon> g')) \<cdot>\<^sub>C G (\<Phi> (G (F g), G (F g')) \<cdot>\<^sub>D
(\<eta> (F g) \<star>\<^sub>D \<eta> (F g')))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (g, g')))"
proof -
have "(\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>C \<epsilon> (G (F g) \<star>\<^sub>C G (F g')) = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C G (F (\<epsilon> g \<star>\<^sub>C \<epsilon> g'))"
using assms \<epsilon>_naturality [of "\<epsilon> g \<star>\<^sub>C \<epsilon> g'"] by simp
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C
G ((F (\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>D \<Phi> (G (F g), G (F g'))) \<cdot>\<^sub>D (\<eta> (F g) \<star>\<^sub>D \<eta> (F g'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (g, g')))"
- using assms C.VV.ide_char C.VV.arr_char FF_def D.comp_assoc by auto
+ using assms C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def D.comp_assoc by auto
also have "... = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C
G (\<Phi> (g, g') \<cdot>\<^sub>D (F (\<epsilon> g) \<star>\<^sub>D F (\<epsilon> g')) \<cdot>\<^sub>D (\<eta> (F g) \<star>\<^sub>D \<eta> (F g'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (g, g')))"
proof -
have "F (\<epsilon> g \<star>\<^sub>C \<epsilon> g') \<cdot>\<^sub>D \<Phi> (G (F g), G (F g')) = \<Phi> (g, g') \<cdot>\<^sub>D (F (\<epsilon> g) \<star>\<^sub>D F (\<epsilon> g'))"
- using assms C.VV.arr_char D.VV.ide_char D.VV.arr_char FF_def
- \<Phi>.naturality [of "(\<epsilon> g, \<epsilon> g')"] C.VV.dom_char C.VV.cod_char
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
+ \<Phi>.naturality [of "(\<epsilon> g, \<epsilon> g')"] C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C
G (\<Phi> (g, g') \<cdot>\<^sub>D (F (\<epsilon> g) \<cdot>\<^sub>D \<eta> (F g) \<star>\<^sub>D F (\<epsilon> g') \<cdot>\<^sub>D \<eta> (F g'))) \<cdot>\<^sub>C
C.inv (G (\<Phi> (g, g')))"
using assms D.interchange by simp
also have "... = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C G (\<Phi> (g, g') \<cdot>\<^sub>D (F g \<star>\<^sub>D F g')) \<cdot>\<^sub>C C.inv (G (\<Phi> (g, g')))"
using assms \<eta>\<epsilon>.triangle_G Fo\<epsilon>_\<eta>oF.map_def [of g] Fo\<epsilon>_\<eta>oF.map_def [of g'] by simp
also have "... = \<epsilon> (g \<star>\<^sub>C g') \<cdot>\<^sub>C G (\<Phi> (g, g')) \<cdot>\<^sub>C C.inv (G (\<Phi> (g, g')))"
using assms D.comp_arr_dom cmp_simps(1) cmp_simps(4) by auto
also have "... = \<epsilon> (g \<star>\<^sub>C g')"
- using assms D.comp_arr_dom C.comp_arr_inv' C.VV.ide_char C.VV.arr_char
+ using assms D.comp_arr_dom C.comp_arr_inv' C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
cmp_components_are_iso C.comp_arr_dom
by simp
finally show ?thesis by simp
qed
lemma G_preserves_hcomp:
assumes "D.hseq \<nu> \<nu>'"
shows "G (\<nu> \<star>\<^sub>D \<nu>') = \<Phi>\<^sub>G.map (D.cod \<nu>, D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>') \<cdot>\<^sub>C
C.inv (\<Phi>\<^sub>G.map (D.dom \<nu>, D.dom \<nu>'))"
proof -
have "G (\<nu> \<star>\<^sub>D \<nu>') \<cdot>\<^sub>C \<Phi>\<^sub>G.map (D.dom \<nu>, D.dom \<nu>') =
\<Phi>\<^sub>G.map (D.cod \<nu>, D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>')"
- using assms \<Phi>\<^sub>G.naturality D.VV.arr_char D.VV.cod_char G.FF_def D.VV.dom_char by auto
+ using assms \<Phi>\<^sub>G.naturality D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.cod_char\<^sub>S\<^sub>b\<^sub>C G.FF_def D.VV.dom_char\<^sub>S\<^sub>b\<^sub>C by auto
moreover have "C.seq (\<Phi>\<^sub>G.map (D.cod \<nu>, D.cod \<nu>')) (G \<nu> \<star>\<^sub>C G \<nu>')"
proof
show "\<guillemotleft>G \<nu> \<star>\<^sub>C G \<nu>' :
G (D.dom \<nu>) \<star>\<^sub>C G (D.dom \<nu>') \<Rightarrow>\<^sub>C G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>')\<guillemotright>"
using assms G_in_hom(2) C.hcomp_in_vhom by auto
show "\<guillemotleft>\<Phi>\<^sub>G.map (D.cod \<nu>, D.cod \<nu>') :
G (D.cod \<nu>) \<star>\<^sub>C G (D.cod \<nu>') \<Rightarrow>\<^sub>C G (D.cod \<nu> \<star>\<^sub>D D.cod \<nu>')\<guillemotright>"
- using assms \<Phi>\<^sub>G\<^sub>0_in_hom \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char by auto
+ using assms \<Phi>\<^sub>G\<^sub>0_in_hom \<Phi>\<^sub>G.map_simp_ide D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
moreover have "C.iso (\<Phi>\<^sub>G.map (D.dom \<nu>, D.dom \<nu>'))"
- using assms \<Phi>\<^sub>G.components_are_iso D.VV.ide_char D.VV.arr_char by auto
+ using assms \<Phi>\<^sub>G.components_are_iso D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show ?thesis
using assms \<Phi>\<^sub>G.components_are_iso C.comp_assoc
C.invert_side_of_triangle(2)
[of "\<Phi>\<^sub>G.map (D.cod \<nu>, D.cod \<nu>') \<cdot>\<^sub>C (G \<nu> \<star>\<^sub>C G \<nu>')"
"G (\<nu> \<star>\<^sub>D \<nu>')" "\<Phi>\<^sub>G.map (D.dom \<nu>, D.dom \<nu>')"]
by simp
qed
lemma coherence_LHS:
assumes "D.ide f" and "D.ide g" and "D.ide h"
and "src\<^sub>D f = trg\<^sub>D g" and "src\<^sub>D g = trg\<^sub>D h"
shows "F (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h))
= (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
\<a>\<^sub>D[F (G f), F (G g), F (G h)] \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "F (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h))
= F (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C
(G (D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)) \<cdot>\<^sub>C
(G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G f, G g))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G g)) \<star>\<^sub>C G h))"
using assms \<Phi>\<^sub>G_map_simp_ide C.comp_assoc by simp
also have "... = F (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C
(G (D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)) \<cdot>\<^sub>C
(G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>C G h) \<cdot>\<^sub>C
(G (D.inv (\<Phi> (G f, G g))) \<star>\<^sub>C G h) \<cdot>\<^sub>C
(C.inv (\<epsilon> (G f \<star>\<^sub>C G g)) \<star>\<^sub>C G h))"
using assms C.whisker_right D.comp_assoc by simp
also have "... = F (G \<a>\<^sub>D[f, g, h]) \<cdot>\<^sub>D
F (G (D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>C G h) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G f, G g))) \<star>\<^sub>C G h) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g)) \<star>\<^sub>C G h)"
proof -
have "C.arr (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C
(G (D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)) \<cdot>\<^sub>C
(G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>C G h) \<cdot>\<^sub>C
(G (D.inv (\<Phi> (G f, G g))) \<star>\<^sub>C G h) \<cdot>\<^sub>C
(C.inv (\<epsilon> (G f \<star>\<^sub>C G g)) \<star>\<^sub>C G h))"
using assms \<Phi>\<^sub>G_simps G.FF_def G\<^sub>0_props by auto
thus ?thesis
by (metis C.seqE as_nat_trans.preserves_comp_2)
qed
also have "... = F (G \<a>\<^sub>D[f, g, h]) \<cdot>\<^sub>D
F (G (D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
(\<Phi> (G (f \<star>\<^sub>D g), G h) \<cdot>\<^sub>D
(F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
((D.inv (\<Phi> (G (F (G f) \<star>\<^sub>D F (G g)), G h))) \<cdot>\<^sub>D
(\<Phi> (G (F (G f) \<star>\<^sub>D F (G g)), G h)) \<cdot>\<^sub>D
(F (G (D.inv (\<Phi> (G f, G g)))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
((D.inv (\<Phi> (G (F (G f \<star>\<^sub>C G g)), G h))) \<cdot>\<^sub>D
(\<Phi> (G (F (G f \<star>\<^sub>C G g)), G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h)))"
using assms G\<^sub>0_props preserves_hcomp FF_def D.comp_assoc by auto
also have "... = F (G \<a>\<^sub>D[f, g, h]) \<cdot>\<^sub>D
F (G (D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
(\<Phi> (G (f \<star>\<^sub>D g), G h) \<cdot>\<^sub>D
(F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(F (G (D.inv (\<Phi> (G f, G g)))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h)))"
proof -
have "((D.inv (\<Phi> (G (F (G f) \<star>\<^sub>D F (G g)), G h))) \<cdot>\<^sub>D
(\<Phi> (G (F (G f) \<star>\<^sub>D F (G g)), G h))) \<cdot>\<^sub>D
(F (G (D.inv (\<Phi> (G f, G g)))) \<star>\<^sub>D F (G h))
= F (G (D.inv (\<Phi> (G f, G g)))) \<star>\<^sub>D F (G h)"
- using assms \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char G\<^sub>0_props FF_def
+ using assms \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C G\<^sub>0_props FF_def
D.comp_inv_arr' D.comp_cod_arr
by simp
moreover have "(D.inv (\<Phi> (G (F (G f \<star>\<^sub>C G g)), G h)) \<cdot>\<^sub>D
\<Phi> (G (F (G f \<star>\<^sub>C G g)), G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) =
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)"
- using assms D.comp_cod_arr \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char
+ using assms D.comp_cod_arr \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
G\<^sub>0_props FF_def D.comp_inv_arr'
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
((D.inv (\<eta> ((f \<star>\<^sub>D g) \<star>\<^sub>D h))) \<cdot>\<^sub>D
(\<eta> ((f \<star>\<^sub>D g) \<star>\<^sub>D h)) \<cdot>\<^sub>D
(D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(((D.inv (\<eta> (F (G (f \<star>\<^sub>D g)) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(\<eta> (F (G (f \<star>\<^sub>D g)) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h))) \<cdot>\<^sub>D
D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
(\<Phi> (G (f \<star>\<^sub>D g), G h) \<cdot>\<^sub>D
((\<eta> (f \<star>\<^sub>D g) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G g)))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
((\<eta> (F (G f) \<star>\<^sub>D F (G g)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f \<star>\<^sub>C G g)))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h)))"
proof -
have "\<And>\<nu>. D.arr \<nu> \<Longrightarrow> F (G \<nu>) = \<eta> (D.cod \<nu>) \<cdot>\<^sub>D \<nu> \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>))"
by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
\<eta>_naturality(1-2) \<eta>_simps(1,6))
thus ?thesis
using assms D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
(D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
(\<Phi> (G (f \<star>\<^sub>D g), G h) \<cdot>\<^sub>D
((\<eta> (f \<star>\<^sub>D g) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(\<eta> (F (G f) \<star>\<^sub>D F (G g)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h)))"
proof -
have "(D.inv (\<eta> ((f \<star>\<^sub>D g) \<star>\<^sub>D h)) \<cdot>\<^sub>D \<eta> ((f \<star>\<^sub>D g) \<star>\<^sub>D h)) \<cdot>\<^sub>D
(D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))
= D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h)"
using assms D.comp_inv_arr' D.comp_cod_arr \<eta>_in_hom by simp
moreover have "((D.inv (\<eta> (F (G (f \<star>\<^sub>D g)) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
\<eta> (F (G (f \<star>\<^sub>D g)) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h))
= D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h))"
using assms D.comp_inv_arr' D.comp_cod_arr by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
((D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D
(\<Phi> (G (f \<star>\<^sub>D g), G h) \<cdot>\<^sub>D
(\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(((D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))))"
proof -
have "(\<eta> (f \<star>\<^sub>D g) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(\<eta> (F (G f) \<star>\<^sub>D F (G g)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))
= (\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(((D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(\<eta> (F (G f) \<star>\<^sub>D F (G g)) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
(D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))"
using assms D.comp_assoc D.whisker_right by simp
also have "... = (\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))"
proof -
have "((D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(\<eta> (F (G f) \<star>\<^sub>D F (G g)) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h))
= D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)"
using assms D.comp_inv_arr' D.comp_cod_arr
D.interchange [of "D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G g)))" "\<eta> (F (G f) \<star>\<^sub>D F (G g))"
"F (G h)" "F (G h)"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
((D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
((D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
(\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "((D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))
= D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "((D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(F (C.inv (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))
= ((D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(D.inv (F (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
using assms by (simp add: preserves_inv)
also have "... = (D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<cdot>\<^sub>D
D.inv (F (\<epsilon> (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h) \<cdot>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
using assms \<Phi>.components_are_iso
D.interchange [of "D.inv (\<eta> (F (G f \<star>\<^sub>C G g)))" "D.inv (F (\<epsilon> (G f \<star>\<^sub>C G g)))"
"F (G h)" "F (G h)"]
by simp
also have "... = (D.inv (F (\<epsilon> (G f \<star>\<^sub>C G g)) \<cdot>\<^sub>D \<eta> (F (G f \<star>\<^sub>C G g))) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "D.iso (F (\<epsilon> (G f \<star>\<^sub>C G g)) \<cdot>\<^sub>D \<eta> (F (G f \<star>\<^sub>C G g)))"
using assms by auto
hence "D.inv (\<eta> (F (G f \<star>\<^sub>C G g))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G f \<star>\<^sub>C G g))) =
D.inv (F (\<epsilon> (G f \<star>\<^sub>C G g)) \<cdot>\<^sub>D \<eta> (F (G f \<star>\<^sub>C G g)))"
using assms by (simp add: D.inv_comp_right(1))
thus ?thesis
using assms by simp
qed
also have "... = (F (G f \<star>\<^sub>C G g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
using assms \<eta>\<epsilon>.triangle_G Fo\<epsilon>_\<eta>oF.map_def [of "G f \<star>\<^sub>C G g"] by simp
also have "... = D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
using assms D.comp_cod_arr by simp
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
((D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
((D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
(\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "((D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))))) \<cdot>\<^sub>D
(\<Phi> (G (f \<star>\<^sub>D g), G h))
= \<Phi> (G (f \<star>\<^sub>D g), G h)"
proof -
have "D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D F (C.inv (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) =
D.inv (\<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)))"
using assms preserves_inv by simp
also have "... = D.inv (F (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)) \<cdot>\<^sub>D \<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)))"
proof -
have "D.iso (F (\<epsilon> (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)) \<cdot>\<^sub>D \<eta> (F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)))"
using assms by auto
thus ?thesis
using assms D.inv_comp_right(1) by simp
qed
also have "... = F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)"
using assms \<eta>\<epsilon>.triangle_G Fo\<epsilon>_\<eta>oF.map_def [of "G (f \<star>\<^sub>D g) \<star>\<^sub>C G h"] by simp
finally show ?thesis
using assms D.comp_cod_arr [of "\<Phi> (G (f \<star>\<^sub>D g), G h)" "F (G (f \<star>\<^sub>D g) \<star>\<^sub>C G h)"]
by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
(((D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
(\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "((D.inv (\<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D \<Phi> (G (f \<star>\<^sub>D g), G h)) \<cdot>\<^sub>D
(\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h)))
= \<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h)"
- using assms D.comp_inv_arr' \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char
+ using assms D.comp_inv_arr' \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
G\<^sub>0_props FF_def D.comp_cod_arr
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(\<a>\<^sub>D[f, g, h] \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
proof -
have "((D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D (\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h))
= (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D D.inv (\<eta> h)"
proof -
have "(D.inv (\<eta> (f \<star>\<^sub>D g)) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D (\<eta> (f \<star>\<^sub>D g) \<star>\<^sub>D F (G h))
= (f \<star>\<^sub>D g) \<star>\<^sub>D D.inv (\<eta> h)"
using assms D.comp_inv_arr' D.comp_arr_dom
D.interchange [of "D.inv (\<eta> (f \<star>\<^sub>D g))" "\<eta> (f \<star>\<^sub>D g)" "D.inv (\<eta> h)" "F (G h)"]
D.invert_side_of_triangle(2)
by simp
moreover have "((f \<star>\<^sub>D g) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D F (G h))
= (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)) \<star>\<^sub>D D.inv (\<eta> h)"
using assms D.comp_cod_arr D.comp_arr_dom
D.interchange [of "f \<star>\<^sub>D g" "D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g)" "D.inv (\<eta> h)" "F (G h)"]
by simp
ultimately show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
\<a>\<^sub>D[F (G f), F (G g), F (G h)]) \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
using assms D.assoc_naturality [of "D.inv (\<eta> f)" "D.inv (\<eta> g)" "D.inv (\<eta> h)"]
D.comp_assoc
by simp
finally show ?thesis
using D.comp_assoc by simp
qed
lemma coherence_RHS:
assumes "D.ide f" and "D.ide g" and "D.ide h"
and "src\<^sub>D f = trg\<^sub>D g" and "src\<^sub>D g = trg\<^sub>D h"
shows "F (\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)))
= (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
proof -
have "F (\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)))
= F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C G (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G g, G h))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))"
using assms \<Phi>\<^sub>G_map_simp_ide C.comp_assoc by simp
also have "... = F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C G (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C G (D.inv (\<Phi> (G g, G h)))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))"
using assms C.whisker_left by simp
also have "... = F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h))))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
F (G f \<star>\<^sub>C G (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
F (G f \<star>\<^sub>C G (D.inv (\<Phi> (G g, G h)))) \<cdot>\<^sub>D
F (G f \<star>\<^sub>C C.inv (\<epsilon> (G g \<star>\<^sub>C G h)))"
proof -
have "C.arr (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>C
G (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>C
C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C G (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C G (D.inv (\<Phi> (G g, G h)))) \<cdot>\<^sub>C
(G f \<star>\<^sub>C C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))"
using assms G\<^sub>0_props by auto
thus ?thesis
using assms by (metis C.seqE as_nat_trans.preserves_comp_2)
qed
also have "... = F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h))))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
(\<Phi> (G f, G (g \<star>\<^sub>D h)) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (G (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)))) \<cdot>\<^sub>D
((D.inv (\<Phi> (G f, G (F (G g) \<star>\<^sub>D F (G h))))) \<cdot>\<^sub>D
(\<Phi> (G f, G (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (G (D.inv (\<Phi> (G g, G h)))))) \<cdot>\<^sub>D
((D.inv (\<Phi> (G f, G (F (G g \<star>\<^sub>C G h))))) \<cdot>\<^sub>D
(\<Phi> (G f, G (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h)))"
using assms preserves_hcomp G\<^sub>0_props D.comp_assoc by simp
also have "... = F (G (D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
F (G (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h))))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
(\<Phi> (G f, G (g \<star>\<^sub>D h)) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (G (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (G (D.inv (\<Phi> (G g, G h))))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h)))"
proof -
have "D.inv (\<Phi> (G f, G (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
\<Phi> (G f, G (F (G g) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (G (D.inv (\<Phi> (G g, G h)))))
= (F (G f) \<star>\<^sub>D F (G (D.inv (\<Phi> (G g, G h)))))"
proof -
have "D.inv (\<Phi> (G f, G (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D \<Phi> (G f, G (F (G g) \<star>\<^sub>D F (G h)))
= F (G f) \<star>\<^sub>D F (G (F (G g) \<star>\<^sub>D F (G h)))"
using assms D.comp_inv_arr D.inv_is_inverse G\<^sub>0_props FF_def by simp
moreover have "... = D.cod (F (G f) \<star>\<^sub>D F (G (D.inv (\<Phi> (G g, G h)))))"
using assms G\<^sub>0_props by simp
moreover have "D.hseq (F (G f)) (F (G (D.inv (\<Phi> (G g, G h)))))"
using assms G\<^sub>0_props by auto
ultimately show ?thesis
using assms D.comp_assoc
D.comp_cod_arr [of "F (G f) \<star>\<^sub>D F (G (D.inv (\<Phi> (G g, G h))))"
"F (G f) \<star>\<^sub>D F (G (F (G g) \<star>\<^sub>D F (G h)))"]
by metis
qed
moreover have "D.inv (\<Phi> (G f, G (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(\<Phi> (G f, G (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))
= (F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))"
proof -
have "D.inv (\<Phi> (G f, G (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D \<Phi> (G f, G (F (G g \<star>\<^sub>C G h)))
= F (G f) \<star>\<^sub>D F (G (F (G g \<star>\<^sub>C G h)))"
using assms D.comp_inv_arr D.inv_is_inverse G\<^sub>0_props by simp
moreover have "... = D.cod (F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))"
using assms by simp
moreover have "D.hseq (F (G f)) (F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))"
using assms by (intro D.hseqI') auto
ultimately show ?thesis
using assms D.comp_assoc
D.comp_cod_arr [of "F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h)))"
"F (G f) \<star>\<^sub>D F (G (F (G g \<star>\<^sub>C G h)))"]
by metis
qed
ultimately show ?thesis by simp
qed
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
((D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h))))) \<cdot>\<^sub>D
\<eta> (F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
((D.inv (\<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))))) \<cdot>\<^sub>D
\<Phi> (G f, G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D
\<eta> (g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D
\<eta> (F (G g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G g, G h)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
proof -
have "\<And>\<nu>. D.arr \<nu> \<Longrightarrow> F (G \<nu>) = \<eta> (D.cod \<nu>) \<cdot>\<^sub>D \<nu> \<cdot>\<^sub>D D.inv (\<eta> (D.dom \<nu>))"
by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
\<eta>_naturality(1-2) \<eta>_simps(1,6))
thus ?thesis
using assms D.comp_assoc by simp
qed
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
\<Phi> (G f, G (g \<star>\<^sub>D h)) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D
\<eta> (g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D
\<eta> (F (G g) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G g, G h)) \<cdot>\<^sub>D
D.inv (\<eta> (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
proof -
have "D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D \<eta> (F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))
= D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))"
proof -
have "D.inv (\<eta> (F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D \<eta> (F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h))) =
F (G f) \<star>\<^sub>D F (G (g \<star>\<^sub>D h))"
using assms D.comp_inv_arr D.inv_is_inverse by simp
moreover have "... = D.cod (D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h))))"
using assms by simp
ultimately show ?thesis
using assms D.comp_cod_arr [of "D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))"]
by (metis D.arr_inv D.comp_assoc D.hcomp_simps(2) D.ideD(1) D.ide_hcomp
G.preserves_ide G_simps(2) G_simps(3) cmp_components_are_iso)
qed
moreover have "(D.inv (\<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))))) \<cdot>\<^sub>D
\<Phi> (G f, G (g \<star>\<^sub>D h))
= \<Phi> (G f, G (g \<star>\<^sub>D h))"
proof -
have "D.inv (\<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))))
= F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))"
proof -
have "D.inv (\<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D F (C.inv (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) =
D.inv (\<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D D.inv (F (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))))"
using assms by (simp add: preserves_inv)
also have "... = D.inv (F (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D \<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))))"
proof -
have "D.iso (F (\<epsilon> (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D \<eta> (F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))))"
using assms by auto
thus ?thesis
using assms by (simp add: D.inv_comp_right(1))
qed
also have "... = F (G f \<star>\<^sub>C G (g \<star>\<^sub>D h))"
using assms Fo\<epsilon>_\<eta>oF.map_def [of "G f \<star>\<^sub>C G (g \<star>\<^sub>D h)"] \<eta>\<epsilon>.triangle_G by simp
finally show ?thesis by blast
qed
moreover have "... = D.cod (\<Phi> (G f, G (g \<star>\<^sub>D h)))"
using assms by simp
moreover have "D.arr (\<Phi> (G f, G (g \<star>\<^sub>D h)))"
using assms by auto
ultimately show ?thesis
using D.comp_cod_arr by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
((D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
\<Phi> (G f, G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)) \<cdot>\<^sub>D
(((F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h)))) \<cdot>\<^sub>D
((F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
using assms D.whisker_left D.comp_assoc by simp
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (g \<star>\<^sub>D h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
proof -
have "(D.inv (\<Phi> (G f, G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D \<Phi> (G f, G (g \<star>\<^sub>D h))) \<cdot>\<^sub>D (F (G f) \<star>\<^sub>D \<eta> (g \<star>\<^sub>D h))
= (F (G f) \<star>\<^sub>D \<eta> (g \<star>\<^sub>D h))"
using assms D.comp_inv_arr' G\<^sub>0_props D.comp_cod_arr by simp
moreover have "((F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h)))
= F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h))"
proof -
have "(F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (F (G g) \<star>\<^sub>D F (G h)))
= F (G f) \<star>\<^sub>D F (G g) \<star>\<^sub>D F (G h)"
proof -
have "(F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (F (G g) \<star>\<^sub>D F (G h)))
= F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g) \<star>\<^sub>D F (G h))) \<cdot>\<^sub>D \<eta> (F (G g) \<star>\<^sub>D F (G h))"
using assms D.interchange [of "F (G f)" "F (G f)"] by simp
also have "... = F (G f) \<star>\<^sub>D F (G g) \<star>\<^sub>D F (G h)"
using assms D.comp_inv_arr' by simp
finally show ?thesis by blast
qed
moreover have "... = D.cod (F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h)))"
using assms by simp
moreover have "D.arr (F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h)))"
using assms by simp
ultimately show ?thesis
using D.comp_cod_arr by simp
qed
moreover have "((F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))
= D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
proof -
have "(F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))
= F (G f) \<star>\<^sub>D F (G g \<star>\<^sub>C G h)"
proof -
have "(F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g \<star>\<^sub>C G h)))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h))))
= F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g \<star>\<^sub>C G h))) \<cdot>\<^sub>D F (C.inv (\<epsilon> (G g \<star>\<^sub>C G h)))"
using assms D.interchange [of "F (G f)" "F (G f)"] by simp
also have "... = F (G f) \<star>\<^sub>D D.inv (\<eta> (F (G g \<star>\<^sub>C G h))) \<cdot>\<^sub>D
D.inv (F (\<epsilon> (G g \<star>\<^sub>C G h)))"
using assms preserves_inv by simp
also have "... = F (G f) \<star>\<^sub>D D.inv (F (\<epsilon> (G g \<star>\<^sub>C G h)) \<cdot>\<^sub>D \<eta> (F (G g \<star>\<^sub>C G h)))"
proof -
have "D.iso (F (\<epsilon> (G g \<star>\<^sub>C G h)) \<cdot>\<^sub>D \<eta> (F (G g \<star>\<^sub>C G h)))"
using assms by auto
thus ?thesis
using assms by (simp add: D.inv_comp_right(1))
qed
also have "... = F (G f) \<star>\<^sub>D F (G g \<star>\<^sub>C G h)"
using assms \<eta>\<epsilon>.triangle_G Fo\<epsilon>_\<eta>oF.map_def [of "G g \<star>\<^sub>C G h"] by simp
finally show ?thesis by blast
qed
thus ?thesis using assms D.comp_cod_arr by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))"
proof -
have "((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))
= D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)"
using assms D.comp_cod_arr D.comp_arr_dom D.interchange
proof -
have "((D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D \<eta> (g \<star>\<^sub>D h))) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))
= (D.inv (\<eta> f) \<cdot>\<^sub>D F (G f) \<star>\<^sub>D D.inv (\<eta> (g \<star>\<^sub>D h)) \<cdot>\<^sub>D \<eta> (g \<star>\<^sub>D h)) \<cdot>\<^sub>D
(F (G f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))"
using assms D.interchange by simp
also have "... = (D.inv (\<eta> f) \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D (F (G f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))"
using assms D.comp_inv_arr' D.comp_arr_dom by simp
also have "... = D.inv (\<eta> f) \<cdot>\<^sub>D F (G f) \<star>\<^sub>D (g \<star>\<^sub>D h) \<cdot>\<^sub>D (D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))"
using assms D.interchange by simp
also have "... = D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h)"
using assms D.comp_arr_dom D.comp_cod_arr by simp
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis by blast
qed
interpretation G: pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G.map
proof
show "\<And>f g h. \<lbrakk>D.ide f; D.ide g; D.ide h; src\<^sub>D f = trg\<^sub>D g; src\<^sub>D g = trg\<^sub>D h\<rbrakk>
\<Longrightarrow> G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h) =
\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)) \<cdot>\<^sub>C \<a>\<^sub>C[G f, G g, G h]"
proof -
fix f g h
assume f: "D.ide f" and g: "D.ide g" and h: "D.ide h"
assume fg: "src\<^sub>D f = trg\<^sub>D g" and gh: "src\<^sub>D g = trg\<^sub>D h"
have "F (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h)) =
F (\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)) \<cdot>\<^sub>C \<a>\<^sub>C[G f, G g, G h])"
proof -
have "F (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h))
= (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
\<a>\<^sub>D[F (G f), F (G g), F (G h)] \<cdot>\<^sub>D
(D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))"
using f g h fg gh coherence_LHS by simp
also have "... = (\<eta> (f \<star>\<^sub>D g \<star>\<^sub>D h) \<cdot>\<^sub>D
(D.inv (\<eta> f) \<star>\<^sub>D D.inv (\<eta> g) \<star>\<^sub>D D.inv (\<eta> h))) \<cdot>\<^sub>D
((F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))) \<cdot>\<^sub>D
F \<a>\<^sub>C[G f, G g, G h]"
proof -
have "\<a>\<^sub>D[F (G f), F (G g), F (G h)] \<cdot>\<^sub>D (D.inv (\<Phi> (G f, G g)) \<star>\<^sub>D F (G h)) \<cdot>\<^sub>D
D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h))
= \<a>\<^sub>D[F (G f), F (G g), F (G h)] \<cdot>\<^sub>D D.inv (\<Phi> (G f \<star>\<^sub>C G g, G h) \<cdot>\<^sub>D
(\<Phi> (G f, G g) \<star>\<^sub>D F (G h)))"
proof -
have "D.iso (\<Phi> (G f \<star>\<^sub>C G g, G h) \<cdot>\<^sub>D (\<Phi> (G f, G g) \<star>\<^sub>D F (G h)))"
- using f g h fg gh \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char FF_def
+ using f g h fg gh \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
by (intro D.iso_hcomp D.isos_compose D.seqI) auto
thus ?thesis
using f g h fg gh
- by (simp add: C.VV.arr_char D.inv_comp_right(1))
+ by (simp add: C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.inv_comp_right(1))
qed
also have "... = D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h) \<cdot>\<^sub>D (F (G f) \<star>\<^sub>D \<Phi> (G g, G h))) \<cdot>\<^sub>D
F \<a>\<^sub>C[G f, G g, G h]"
using f g h fg gh cmp_simps assoc_coherence D.comp_assoc D.isos_compose
- \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
D.invert_opposite_sides_of_square
by simp
also have "... = ((F (G f) \<star>\<^sub>D D.inv (\<Phi> (G g, G h))) \<cdot>\<^sub>D
D.inv (\<Phi> (G f, G g \<star>\<^sub>C G h))) \<cdot>\<^sub>D
F \<a>\<^sub>C[G f, G g, G h]"
proof -
have "D.iso (\<Phi> (G f, G g \<star>\<^sub>C G h) \<cdot>\<^sub>D (F (G f) \<star>\<^sub>D \<Phi> (G g, G h)))"
- using f g h fg gh \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char FF_def
+ using f g h fg gh \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
by auto
thus ?thesis
- using f g h fg gh \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char FF_def
- by (simp add: C.VV.arr_char D.comp_assoc D.inv_comp_right(1))
+ using f g h fg gh \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
+ by (simp add: C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.comp_assoc D.inv_comp_right(1))
qed
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = F (\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h))) \<cdot>\<^sub>D
F \<a>\<^sub>C[G f, G g, G h]"
using f g h fg gh coherence_RHS D.comp_assoc by simp
also have "... = F ((\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h))) \<cdot>\<^sub>C
\<a>\<^sub>C[G f, G g, G h])"
using f g h fg gh \<Phi>\<^sub>G_simps by auto
also have "... = F (\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)) \<cdot>\<^sub>C
\<a>\<^sub>C[G f, G g, G h])"
using C.comp_assoc by simp
finally show ?thesis by simp
qed
moreover have "C.par (G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C
(\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h))
(\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)) \<cdot>\<^sub>C \<a>\<^sub>C[G f, G g, G h])"
using f g h fg gh \<Phi>\<^sub>G_simps by auto
ultimately show "G \<a>\<^sub>D[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>G.map (f \<star>\<^sub>D g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>G.map (f, g) \<star>\<^sub>C G h) =
\<Phi>\<^sub>G.map (f, g \<star>\<^sub>D h) \<cdot>\<^sub>C (G f \<star>\<^sub>C \<Phi>\<^sub>G.map (g, h)) \<cdot>\<^sub>C \<a>\<^sub>C[G f, G g, G h]"
using is_faithful by blast
qed
qed
interpretation GF: composite_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi> G \<Phi>\<^sub>G.map
..
interpretation FG: composite_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G.map F \<Phi>
..
interpretation I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C ..
interpretation I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
lemma \<eta>_equals_FG_unit:
assumes "D.obj a"
shows "\<eta> a = FG.unit a"
proof (intro FG.unit_eqI)
show "D.obj a" by fact
show "D.iso (\<eta> a)"
using assms by auto
show "\<guillemotleft>\<eta> a : FG.map\<^sub>0 a \<Rightarrow>\<^sub>D FG.map a\<guillemotright>"
using assms \<eta>_in_hom [of a] FG.map\<^sub>0_def G\<^sub>0_props D.obj_def
by (metis D.ideD(2) D.ideD(3) D.objE D.vconn_implies_hpar(3) \<eta>.preserves_cod
\<eta>_simps(5))
show "\<eta> a \<cdot>\<^sub>D \<i>\<^sub>D[FG.map\<^sub>0 a] = (FG.map \<i>\<^sub>D[a] \<cdot>\<^sub>D FG.cmp (a, a)) \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D \<eta> a)"
proof -
have "(FG.map \<i>\<^sub>D[a] \<cdot>\<^sub>D FG.cmp (a, a)) \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D \<eta> a) =
F (G \<i>\<^sub>D[a]) \<cdot>\<^sub>D FG.cmp (a, a) \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D \<eta> a)"
using assms D.comp_assoc by simp
also have "... = F (G \<i>\<^sub>D[a]) \<cdot>\<^sub>D F (G (I\<^sub>D.cmp (a, a))) \<cdot>\<^sub>D F (\<Phi>\<^sub>G.map (a, a)) \<cdot>\<^sub>D
\<Phi> (G a, G a) \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D \<eta> a)"
- using assms FG.cmp_def D.comp_assoc D.VV.arr_char D.VV.dom_char D.VV.cod_char
+ using assms FG.cmp_def D.comp_assoc D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_char\<^sub>S\<^sub>b\<^sub>C D.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by auto
also have "... = F (G \<i>\<^sub>D[a]) \<cdot>\<^sub>D F (G (I\<^sub>D.cmp (a, a))) \<cdot>\<^sub>D \<eta> (a \<star>\<^sub>D a)"
using assms \<eta>_hcomp by auto
also have "... = F (G \<i>\<^sub>D[a]) \<cdot>\<^sub>D \<eta> (a \<star>\<^sub>D a)"
using assms D.comp_cod_arr by auto
also have "... = \<eta> a \<cdot>\<^sub>D \<i>\<^sub>D[a]"
using assms \<eta>_naturality [of "\<i>\<^sub>D[a]"] by simp
also have "... = \<eta> a \<cdot>\<^sub>D \<i>\<^sub>D[FG.map\<^sub>0 a]"
using assms \<open>\<guillemotleft>\<eta> a : FG.map\<^sub>0 a \<Rightarrow>\<^sub>D FG.map a\<guillemotright>\<close> by fastforce
finally show ?thesis by simp
qed
qed
lemma \<epsilon>_hcomp':
assumes "C.ide g" and "C.ide f" and "src\<^sub>C g = trg\<^sub>C f"
shows "\<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C GF.cmp (g, f) = \<epsilon> g \<star>\<^sub>C \<epsilon> f"
proof -
have "\<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C GF.cmp (g, f)
= (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C C.inv (\<Phi>\<^sub>G.map (F g, F f)) \<cdot>\<^sub>C C.inv (G (\<Phi> (g, f))) \<cdot>\<^sub>C
G (F (g \<star>\<^sub>C f)) \<cdot>\<^sub>C G (\<Phi> (g, f)) \<cdot>\<^sub>C \<Phi>\<^sub>G.map (F g, F f)"
- using assms \<epsilon>_hcomp GF.cmp_def C.VV.arr_char C.comp_cod_arr
- C.comp_inv_arr' C.comp_assoc C.VV.dom_char C.VV.cod_char
+ using assms \<epsilon>_hcomp GF.cmp_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.comp_cod_arr
+ C.comp_inv_arr' C.comp_assoc C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C C.inv (\<Phi>\<^sub>G.map (F g, F f)) \<cdot>\<^sub>C (C.inv (G (\<Phi> (g, f))) \<cdot>\<^sub>C
G (F (g \<star>\<^sub>C f)) \<cdot>\<^sub>C G (\<Phi> (g, f))) \<cdot>\<^sub>C \<Phi>\<^sub>G.map (F g, F f)"
using C.comp_assoc by simp
also have "... = (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C C.inv (\<Phi>\<^sub>G.map (F g, F f)) \<cdot>\<^sub>C (C.inv (G (\<Phi> (g, f))) \<cdot>\<^sub>C
G (\<Phi> (g, f))) \<cdot>\<^sub>C \<Phi>\<^sub>G.map (F g, F f)"
using assms C.comp_ide_arr [of "G (F (g \<star>\<^sub>C f))" "G (\<Phi> (g, f))"]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C C.inv (\<Phi>\<^sub>G.map (F g, F f)) \<cdot>\<^sub>C \<Phi>\<^sub>G.map (F g, F f)"
proof -
have "C.inv (G (\<Phi> (g, f))) \<cdot>\<^sub>C G (\<Phi> (g, f)) = G (F g \<star>\<^sub>D F f)"
using assms C.comp_inv_arr' cmp_components_are_iso C.inv_is_inverse
- C.VV.arr_char C.VV.ide_char cmp_simps(4)
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C cmp_simps(4)
by auto
moreover have "... = C.cod (\<Phi>\<^sub>G.map (F g, F f))"
using assms by simp
ultimately have "(C.inv (G (\<Phi> (g, f))) \<cdot>\<^sub>C G (\<Phi> (g, f))) \<cdot>\<^sub>C \<Phi>\<^sub>G.map (F g, F f) =
\<Phi>\<^sub>G.map (F g, F f)"
using assms C.comp_cod_arr [of "\<Phi>\<^sub>G.map (F g, F f)" "G (F g \<star>\<^sub>D F f)"]
C.ideD(1) G.cmp_simps(1) preserves_ide preserves_src preserves_trg
by presburger
thus ?thesis by simp
qed
also have "... = \<epsilon> g \<star>\<^sub>C \<epsilon> f"
using assms C.comp_inv_arr' C.comp_arr_dom [of "\<epsilon> g \<star>\<^sub>C \<epsilon> f" "G (F g) \<star>\<^sub>C G (F f)"]
by simp
finally show ?thesis by simp
qed
lemma \<epsilon>_inverts_GF_unit:
assumes "C.obj a"
shows "\<epsilon> a \<cdot>\<^sub>C GF.unit a = a"
proof -
have "\<epsilon> a \<cdot>\<^sub>C GF.unit a = I\<^sub>C.unit a"
proof (intro I\<^sub>C.unit_eqI)
show "C.obj a" by fact
show 1: "\<guillemotleft>\<epsilon> a \<cdot>\<^sub>C GF.unit a : I\<^sub>C.map\<^sub>0 a \<Rightarrow>\<^sub>C I\<^sub>C.map a\<guillemotright>"
proof -
have "src\<^sub>C (G (F a)) = src\<^sub>C (I\<^sub>C.map a)"
using assms G\<^sub>0_props C.obj_def' by simp
thus ?thesis
using assms I\<^sub>C.map\<^sub>0_def GF.map\<^sub>0_def GF.unit_in_hom
by (intro C.comp_in_homI') auto
qed
show "C.iso (\<epsilon> a \<cdot>\<^sub>C GF.unit a)"
using assms \<epsilon>_in_hom GF.unit_char(2)
by (intro C.isos_compose) auto
show "(\<epsilon> a \<cdot>\<^sub>C GF.unit a) \<cdot>\<^sub>C \<i>\<^sub>C[I\<^sub>C.map\<^sub>0 a]
= (I\<^sub>C.map \<i>\<^sub>C[a] \<cdot>\<^sub>C I\<^sub>C.cmp (a, a)) \<cdot>\<^sub>C (\<epsilon> a \<cdot>\<^sub>C GF.unit a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a)"
proof -
have "(I\<^sub>C.map \<i>\<^sub>C[a] \<cdot>\<^sub>C I\<^sub>C.cmp (a, a)) \<cdot>\<^sub>C (\<epsilon> a \<cdot>\<^sub>C GF.unit a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a) =
\<i>\<^sub>C[a] \<cdot>\<^sub>C (a \<star>\<^sub>C a) \<cdot>\<^sub>C (\<epsilon> a \<cdot>\<^sub>C GF.unit a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a)"
using assms C.comp_assoc by simp
also have "... = \<i>\<^sub>C[a] \<cdot>\<^sub>C (\<epsilon> a \<cdot>\<^sub>C GF.unit a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a)"
proof -
have "C.hseq (\<epsilon> a \<cdot>\<^sub>C GF.unit a) (\<epsilon> a \<cdot>\<^sub>C GF.unit a)"
using assms GF.unit_simps C.iso_is_arr \<open>C.iso (\<epsilon> a \<cdot>\<^sub>C GF.unit a)\<close>
by (intro C.hseqI') auto
moreover have "a \<star>\<^sub>C a = C.cod (\<epsilon> a \<cdot>\<^sub>C GF.unit a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a)"
proof -
have "C.cod (\<epsilon> a \<cdot>\<^sub>C GF.unit a) = a"
using assms 1 C.obj_simps by auto
moreover have "C.hseq (\<epsilon> a \<cdot>\<^sub>C GF.unit a) (\<epsilon> a \<cdot>\<^sub>C GF.unit a)"
using assms 1 C.src_dom [of "\<epsilon> a \<cdot>\<^sub>C GF.unit a"] C.trg_dom [of "\<epsilon> a \<cdot>\<^sub>C GF.unit a"]
apply (intro C.hseqI')
by auto force
ultimately show ?thesis by auto
qed
ultimately show ?thesis
using C.comp_cod_arr by simp
qed
also have "... = \<i>\<^sub>C[a] \<cdot>\<^sub>C (\<epsilon> a \<star>\<^sub>C \<epsilon> a) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C GF.unit a)"
using assms C.interchange [of "\<epsilon> a" "GF.unit a" "\<epsilon> a" "GF.unit a"]
by (simp add: C.iso_is_arr \<open>C.iso (\<epsilon> a \<cdot>\<^sub>C GF.unit a)\<close>)
also have "... = \<i>\<^sub>C[a] \<cdot>\<^sub>C (\<epsilon> (a \<star>\<^sub>C a) \<cdot>\<^sub>C GF.cmp (a, a)) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C GF.unit a)"
using assms \<epsilon>_hcomp' [of a a] by auto
also have "... = (\<i>\<^sub>C[a] \<cdot>\<^sub>C \<epsilon> (a \<star>\<^sub>C a)) \<cdot>\<^sub>C GF.cmp (a, a) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C GF.unit a)"
using C.comp_assoc by simp
also have "... = (\<epsilon> a \<cdot>\<^sub>C G (F \<i>\<^sub>C[a])) \<cdot>\<^sub>C GF.cmp (a, a) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C GF.unit a)"
using assms \<epsilon>_naturality [of "\<i>\<^sub>C[a]"] by simp
also have "... = \<epsilon> a \<cdot>\<^sub>C (G (F \<i>\<^sub>C[a]) \<cdot>\<^sub>C GF.cmp (a, a)) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C GF.unit a)"
using C.comp_assoc by simp
also have "... = \<epsilon> a \<cdot>\<^sub>C GF.unit a \<cdot>\<^sub>C \<i>\<^sub>C[GF.map\<^sub>0 a]"
using assms GF.unit_char [of a] by simp
also have "... = (\<epsilon> a \<cdot>\<^sub>C GF.unit a) \<cdot>\<^sub>C \<i>\<^sub>C[I\<^sub>C.map\<^sub>0 a]"
proof -
have "GF.map\<^sub>0 a = I\<^sub>C.map\<^sub>0 a"
using assms G\<^sub>0_props(2) [of a] GF.map\<^sub>0_def by auto
thus ?thesis
using assms GF.unit_char [of a] C.comp_assoc by simp
qed
finally show ?thesis
using C.comp_assoc by simp
qed
qed
thus ?thesis
using assms I\<^sub>C.unit_char' by simp
qed
lemma \<eta>_respects_comp:
assumes "D.ide f" and "D.ide g" and "src\<^sub>D g = trg\<^sub>D f"
shows "(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D \<eta> (g \<star>\<^sub>D f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]) \<cdot>\<^sub>D ((g \<star>\<^sub>D f) \<star>\<^sub>D src\<^sub>D f)
= (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
and "(trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<cdot>\<^sub>D \<eta> g \<cdot>\<^sub>D \<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)] \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)] \<cdot>\<^sub>D \<eta> f \<cdot>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D \<a>\<^sub>D[g, f, src\<^sub>D f]
= (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
show "(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D \<eta> (g \<star>\<^sub>D f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]) \<cdot>\<^sub>D ((g \<star>\<^sub>D f) \<star>\<^sub>D src\<^sub>D f)
= (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D \<eta> (g \<star>\<^sub>D f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]) \<cdot>\<^sub>D ((g \<star>\<^sub>D f) \<star>\<^sub>D src\<^sub>D f)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D \<eta> (g \<star>\<^sub>D f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
using assms D.comp_assoc D.comp_arr_dom by simp
also have 1: "... = (\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D F (\<Phi>\<^sub>G.map (g, f))) \<cdot>\<^sub>D
\<Phi> (G g, G f) \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
using assms \<eta>_hcomp D.comp_assoc by simp
also have "... = (trg\<^sub>D g \<star>\<^sub>D F (\<Phi>\<^sub>G.map (g, f))) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g \<star>\<^sub>C G f)] \<cdot>\<^sub>D
\<Phi> (G g, G f)) \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D F (\<Phi>\<^sub>G.map (g, f)) =
(trg\<^sub>D g \<star>\<^sub>D F (\<Phi>\<^sub>G.map (g, f))) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g \<star>\<^sub>C G f)]"
using assms G\<^sub>0_props D.lunit'_naturality [of "F (\<Phi>\<^sub>G.map (g, f))"] \<Phi>\<^sub>G_in_hom [of g f]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((trg\<^sub>D g \<star>\<^sub>D F (\<Phi>\<^sub>G.map (g, f))) \<cdot>\<^sub>D (trg\<^sub>D g \<star>\<^sub>D \<Phi> (G g, G f))) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g \<star>\<^sub>C G f)] \<cdot>\<^sub>D \<Phi> (G g, G f)
= (trg\<^sub>D g \<star>\<^sub>D \<Phi> (G g, G f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)]"
using assms D.lunit'_naturality [of "\<Phi> (G g, G f)"] G\<^sub>0_props by fastforce
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg\<^sub>D g \<star>\<^sub>D F (\<Phi>\<^sub>G.map (g, f)) \<cdot>\<^sub>D \<Phi> (G g, G f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
using assms 1 D.whisker_left [of "trg\<^sub>D g" "F (\<Phi>\<^sub>G.map (g, f))" "\<Phi> (G g, G f)"]
by force
also have "... = (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D (\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have 1: "FG.cmp (g, f) = (F (G (g \<star>\<^sub>D f)) \<cdot>\<^sub>D F (\<Phi>\<^sub>G.map (g, f))) \<cdot>\<^sub>D \<Phi> (G g, G f)"
- using assms FG.cmp_def D.VV.arr_char D.VV.dom_char D.comp_assoc by simp
+ using assms FG.cmp_def D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_char\<^sub>S\<^sub>b\<^sub>C D.comp_assoc by simp
also have "... = F (\<Phi>\<^sub>G.map (g, f)) \<cdot>\<^sub>D \<Phi> (G g, G f)"
proof -
have "D.cod (F (\<Phi>\<^sub>G (g, f))) = F (G (g \<star>\<^sub>D f))"
using assms 1
by (metis (mono_tags, lifting) D.cod_eqI D.ideD(1) D.ide_hcomp D.seqE
FG.cmp_simps'(1) G.preserves_ide preserves_ide)
thus ?thesis
using assms D.comp_cod_arr [of "F (\<Phi>\<^sub>G.map (g, f))" "F (G (g \<star>\<^sub>D f))"]
by fastforce
qed
finally have "FG.cmp (g, f) = F (\<Phi>\<^sub>G.map (g, f)) \<cdot>\<^sub>D \<Phi> (G g, G f)" by simp
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
show "(trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<cdot>\<^sub>D \<eta> g \<cdot>\<^sub>D \<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)] \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)] \<cdot>\<^sub>D \<eta> f \<cdot>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D
\<a>\<^sub>D[g, f, src\<^sub>D f]
= (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D
(\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "(trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<cdot>\<^sub>D \<eta> g \<cdot>\<^sub>D \<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)] \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)] \<cdot>\<^sub>D \<eta> f \<cdot>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D
\<a>\<^sub>D[g, f, src\<^sub>D f]
= (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
(\<eta> g \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
(\<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)] \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)]) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D
\<a>\<^sub>D[g, f, src\<^sub>D f]"
using assms D.comp_assoc D.whisker_right D.whisker_left by simp
also have "... = (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
(\<a>\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<star>\<^sub>D F (G f))) \<cdot>\<^sub>D
(\<eta> g \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
(\<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)] \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)]) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
((g \<star>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D
\<a>\<^sub>D[g, f, src\<^sub>D f])"
- using assms D.comp_cod_arr D.VV.ide_char D.VV.arr_char D.VV.dom_char
+ using assms D.comp_cod_arr D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.dom_char\<^sub>S\<^sub>b\<^sub>C
FG.FF_def G\<^sub>0_props D.comp_assoc
by presburger
also have "... = (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D
(\<eta> g \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
((\<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)]) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)]) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "(g \<star>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D \<a>\<^sub>D[g, f, src\<^sub>D f] = \<r>\<^sub>D[g \<star>\<^sub>D f]"
using assms D.runit_hcomp by simp
moreover have "\<a>\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<star>\<^sub>D F (G f)) =
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)]"
using assms D.lunit_hcomp [of "F (G g)" "F (G f)"] G\<^sub>0_props by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D
(\<eta> g \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
(((g \<star>\<^sub>D \<l>\<^sub>D[F (G f)]) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)])) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<eta> f)) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "(\<r>\<^sub>D[g] \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D D.inv \<a>\<^sub>D[g, src\<^sub>D g, F (G f)] = g \<star>\<^sub>D \<l>\<^sub>D[F (G f)]"
using assms D.triangle' [of g "F (G f)"] G\<^sub>0_props by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D
((\<eta> g \<star>\<^sub>D F (G f)) \<cdot>\<^sub>D
(g \<star>\<^sub>D \<eta> f)) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]"
proof -
have "((g \<star>\<^sub>D \<l>\<^sub>D[F (G f)]) \<cdot>\<^sub>D (g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)])) \<cdot>\<^sub>D (g \<star>\<^sub>D \<eta> f) = g \<star>\<^sub>D \<eta> f"
using assms D.interchange [of g g "\<l>\<^sub>D[F (G f)]" "\<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)]"]
D.comp_arr_inv' D.comp_cod_arr
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g) \<star>\<^sub>D F (G f)] \<cdot>\<^sub>D
(\<eta> g \<star>\<^sub>D \<eta> f) \<cdot>\<^sub>D \<r>\<^sub>D[g \<star>\<^sub>D f]"
using assms D.interchange [of "\<eta> g" g "F (G f)" "\<eta> f"] D.comp_arr_dom D.comp_cod_arr
by simp
finally show ?thesis by simp
qed
qed
lemma \<eta>_respects_unit:
assumes "D.obj a"
shows "(a \<star>\<^sub>D FG.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[a] \<cdot>\<^sub>D \<l>\<^sub>D[a] =
(\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod a)] \<cdot>\<^sub>D \<eta> a \<cdot>\<^sub>D \<r>\<^sub>D[D.dom a]) \<cdot>\<^sub>D (I\<^sub>D.unit a \<star>\<^sub>D a)"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod a)] \<cdot>\<^sub>D \<eta> a \<cdot>\<^sub>D \<r>\<^sub>D[D.dom a]) \<cdot>\<^sub>D (I\<^sub>D.unit a \<star>\<^sub>D a) =
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G a)] \<cdot>\<^sub>D \<eta> a) \<cdot>\<^sub>D \<r>\<^sub>D[a]"
using assms I\<^sub>D.lunit_coherence I\<^sub>D.unit_char' D.comp_arr_dom D.comp_assoc by auto
also have "... = ((a \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[a]) \<cdot>\<^sub>D \<r>\<^sub>D[a]"
using assms D.lunit'_naturality [of "\<eta> a"] by auto
also have "... = (a \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[a] \<cdot>\<^sub>D \<r>\<^sub>D[a]"
using D.comp_assoc by simp
also have "... = (a \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[a] \<cdot>\<^sub>D \<l>\<^sub>D[a]"
using assms D.unitor_coincidence by simp
also have "... = (a \<star>\<^sub>D FG.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[a] \<cdot>\<^sub>D \<l>\<^sub>D[a]"
using assms \<eta>_equals_FG_unit by simp
finally show ?thesis by simp
qed
lemma \<epsilon>_respects_comp:
assumes "C.ide f" and "C.ide g" and "src\<^sub>C g = trg\<^sub>C f"
shows "(trg\<^sub>C g \<star>\<^sub>C g \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C[trg\<^sub>C g, g, f] \<cdot>\<^sub>C (\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<epsilon> g \<cdot>\<^sub>C \<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f] \<cdot>\<^sub>C (G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[G (F g), G (F f), src\<^sub>C f]
= \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
and "(\<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C \<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (g \<star>\<^sub>C f))]) \<cdot>\<^sub>C (GF.cmp (g, f) \<star>\<^sub>C src\<^sub>C f)
= \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
proof -
have "(trg\<^sub>C g \<star>\<^sub>C g \<star>\<^sub>C f) \<cdot>\<^sub>C
\<a>\<^sub>C[trg\<^sub>C g, g, f] \<cdot>\<^sub>C
(\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<epsilon> g \<cdot>\<^sub>C \<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f] \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[G (F g), G (F f), src\<^sub>C f]
= ((trg\<^sub>C g \<star>\<^sub>C g \<star>\<^sub>C f) \<cdot>\<^sub>C
\<a>\<^sub>C[trg\<^sub>C g, g, f]) \<cdot>\<^sub>C
(\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<epsilon> g \<cdot>\<^sub>C \<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f] \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[G (F g), G (F f), src\<^sub>C f]"
using assms C.comp_assoc by simp
also have "... = \<a>\<^sub>C[trg\<^sub>C g, g, f] \<cdot>\<^sub>C
(\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<epsilon> g \<cdot>\<^sub>C \<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f] \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[G (F g), G (F f), src\<^sub>C f]"
using assms C.comp_cod_arr by simp
also have "... = (\<a>\<^sub>C[trg\<^sub>C g, g, f] \<cdot>\<^sub>C
(\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<star>\<^sub>C f)) \<cdot>\<^sub>C
(\<epsilon> g \<star>\<^sub>C f) \<cdot>\<^sub>C
(\<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f] \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f]) \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C
((G (F g) \<star>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[G (F g), G (F f), src\<^sub>C f])"
using assms C.whisker_left C.whisker_right C.comp_assoc by simp
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C
(\<epsilon> g \<star>\<^sub>C f) \<cdot>\<^sub>C
((\<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f]) \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f]) \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C
\<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
using assms G\<^sub>0_props C.lunit_hcomp C.runit_hcomp C.comp_assoc by simp
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C
(\<epsilon> g \<star>\<^sub>C f) \<cdot>\<^sub>C
(((G (F g) \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f])) \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<epsilon> f)) \<cdot>\<^sub>C
\<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
using assms G\<^sub>0_props C.triangle' C.comp_assoc by simp
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C
((\<epsilon> g \<star>\<^sub>C f) \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<epsilon> f)) \<cdot>\<^sub>C
\<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
proof -
have "(G (F g) \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>C (G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f]) = G (F g) \<star>\<^sub>C f"
using assms C.whisker_left [of "G (F g)" "\<l>\<^sub>C[f]" "\<l>\<^sub>C\<^sup>-\<^sup>1[f]"] C.comp_arr_inv'
by simp
moreover have "... = C.cod (G (F g) \<star>\<^sub>C \<epsilon> f)"
using assms G\<^sub>0_props by auto
moreover have "C.hseq (G (F g)) (\<epsilon> f)"
using assms G\<^sub>0_props by simp
ultimately have "((G (F g) \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>C (G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f])) \<cdot>\<^sub>C (G (F g) \<star>\<^sub>C \<epsilon> f) =
(G (F g) \<star>\<^sub>C \<epsilon> f)"
using assms G\<^sub>0_props C.comp_cod_arr by presburger
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
using assms C.interchange [of "\<epsilon> g" "G (F g)" f "\<epsilon> f"] C.comp_cod_arr C.comp_arr_dom
by simp
finally show "(trg\<^sub>C g \<star>\<^sub>C g \<star>\<^sub>C f) \<cdot>\<^sub>C
\<a>\<^sub>C[trg\<^sub>C g, g, f] \<cdot>\<^sub>C
(\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<epsilon> g \<cdot>\<^sub>C \<r>\<^sub>C[G (F g)] \<star>\<^sub>C f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[G (F g), src\<^sub>C g, f] \<cdot>\<^sub>C
(G (F g) \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[G (F g), G (F f), src\<^sub>C f]
= \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
by simp
show "(\<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C \<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (g \<star>\<^sub>C f))]) \<cdot>\<^sub>C (GF.cmp (g, f) \<star>\<^sub>C src\<^sub>C f)
= ..."
proof -
have "(\<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C \<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (g \<star>\<^sub>C f))]) \<cdot>\<^sub>C (GF.cmp (g, f) \<star>\<^sub>C src\<^sub>C f) =
\<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C \<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (g \<star>\<^sub>C f))] \<cdot>\<^sub>C (GF.cmp (g, f) \<star>\<^sub>C src\<^sub>C f)"
using assms C.comp_assoc by simp
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C (\<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C GF.cmp (g, f)) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
proof -
have "src\<^sub>C (GF.cmp (g, f)) = src\<^sub>C f"
using assms G\<^sub>0_props by simp
hence "\<r>\<^sub>C[G (F (g \<star>\<^sub>C f))] \<cdot>\<^sub>C (GF.cmp (g, f) \<star>\<^sub>C src\<^sub>C f) =
GF.cmp (g, f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
- using assms C.runit_naturality [of "GF.cmp (g, f)"] C.VV.arr_char C.VV.cod_char
+ using assms C.runit_naturality [of "GF.cmp (g, f)"] C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
GF.cmp_simps'(1,4-5)
by simp
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C (\<epsilon> g \<star>\<^sub>C \<epsilon> f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F g) \<star>\<^sub>C G (F f)]"
using assms \<epsilon>_hcomp' by simp
ultimately show ?thesis by simp
qed
qed
lemma \<epsilon>_respects_unit:
assumes "C.obj a"
shows "(a \<star>\<^sub>C I\<^sub>C.unit a) \<cdot>\<^sub>C \<r>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<l>\<^sub>C[a] =
(\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod a] \<cdot>\<^sub>C \<epsilon> a \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom a)]) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C a)"
proof -
have "(\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod a] \<cdot>\<^sub>C \<epsilon> a \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom a)]) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C a) =
\<l>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<epsilon> a \<cdot>\<^sub>C \<r>\<^sub>C[G (F a)] \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C a)"
using assms C.comp_assoc by auto
also have "... = (\<l>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<epsilon> a) \<cdot>\<^sub>C GF.unit a \<cdot>\<^sub>C \<r>\<^sub>C[a]"
proof -
have "src\<^sub>C (GF.unit a) = a"
using assms GF.unit_simps(2) GF.map\<^sub>0_def [of a] G\<^sub>0_props
by (simp add: C.obj_simps(1-2))
thus ?thesis
using assms C.runit_naturality [of "GF.unit a"] C.comp_assoc by simp
qed
also have "... = (a \<star>\<^sub>C \<epsilon> a) \<cdot>\<^sub>C (\<l>\<^sub>C\<^sup>-\<^sup>1[G (F a)] \<cdot>\<^sub>C GF.unit a) \<cdot>\<^sub>C \<r>\<^sub>C[a]"
proof -
have "\<l>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<epsilon> a = (a \<star>\<^sub>C \<epsilon> a) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[G (F a)]"
using assms C.lunit'_naturality [of "\<epsilon> a"] by auto
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = ((a \<star>\<^sub>C \<epsilon> a) \<cdot>\<^sub>C (a \<star>\<^sub>C GF.unit a)) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<r>\<^sub>C[a]"
proof -
have "\<l>\<^sub>C\<^sup>-\<^sup>1[G (F a)] \<cdot>\<^sub>C GF.unit a = (a \<star>\<^sub>C GF.unit a) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[a]"
using assms C.lunit'_naturality [of "GF.unit a"] G\<^sub>0_props C.obj_simps
by (simp add: GF.map\<^sub>0_def)
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = (a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<r>\<^sub>C[a]"
using assms C.interchange [of a a "\<epsilon> a" "GF.unit a"] by force
also have "... = (a \<star>\<^sub>C \<epsilon> a \<cdot>\<^sub>C GF.unit a) \<cdot>\<^sub>C \<r>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<l>\<^sub>C[a]"
using assms C.unitor_coincidence by simp
also have "... = (a \<star>\<^sub>C I\<^sub>C.unit a) \<cdot>\<^sub>C \<r>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<l>\<^sub>C[a]"
using assms \<epsilon>_inverts_GF_unit I\<^sub>C.unit_char' by simp
finally show ?thesis by simp
qed
(*
* The following clash with definitions made later in the context of
* equivalence_pseudofunctor. The clash is only flagged when -o export_theory
* is invoked with "isabelle_build", as was pointed out by Fabian Huch in an email
* of 3/31/2022. The reason for the clash is not clear.
* I have distinguished the constants here with a prime because it is unlikely
* that they will be useful outside of the present theory, so the name is not
* so important.
*)
abbreviation counit\<^sub>0'
where "counit\<^sub>0' \<equiv> \<lambda>b. b"
abbreviation counit\<^sub>1'
where "counit\<^sub>1' \<equiv> \<lambda>g. \<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<cdot>\<^sub>D \<eta> g \<cdot>\<^sub>D \<r>\<^sub>D[g]"
interpretation \<epsilon>: pseudonatural_equivalence
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
FG.map FG.cmp I\<^sub>D.map I\<^sub>D.cmp counit\<^sub>0' counit\<^sub>1'
proof
show "\<And>a. D.obj a \<Longrightarrow> D.ide a"
by auto
show "\<And>f. D.ide f \<Longrightarrow> D.iso (\<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)] \<cdot>\<^sub>D \<eta> f \<cdot>\<^sub>D \<r>\<^sub>D[f])"
using D.iso_lunit' D.iso_runit \<eta>_in_hom
by (intro D.isos_compose D.seqI) auto
show "\<And>a. D.obj a \<Longrightarrow> \<guillemotleft>a : src\<^sub>D (FG.map a) \<rightarrow>\<^sub>D src\<^sub>D (I\<^sub>D.map a)\<guillemotright>"
using D.obj_def G\<^sub>0_props(1) by auto
show "\<And>f. D.ide f \<Longrightarrow> \<guillemotleft>\<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)] \<cdot>\<^sub>D \<eta> f \<cdot>\<^sub>D \<r>\<^sub>D[f] :
I\<^sub>D.map f \<star>\<^sub>D src\<^sub>D f \<Rightarrow>\<^sub>D trg\<^sub>D f \<star>\<^sub>D FG.map f\<guillemotright>"
using G\<^sub>0_props
by (intro D.comp_in_homI') auto
show "\<And>\<mu>. D.arr \<mu> \<Longrightarrow>
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (D.cod \<mu>))] \<cdot>\<^sub>D \<eta> (D.cod \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.cod \<mu>]) \<cdot>\<^sub>D (I\<^sub>D.map \<mu> \<star>\<^sub>D src\<^sub>D \<mu>)
= (trg\<^sub>D \<mu> \<star>\<^sub>D FG.map \<mu>) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G (D.dom \<mu>))] \<cdot>\<^sub>D \<eta> (D.dom \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
proof -
fix \<mu>
assume \<mu>: "D.arr \<mu>"
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D \<eta> (D.cod \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.cod \<mu>]) \<cdot>\<^sub>D
(I\<^sub>D.map \<mu> \<star>\<^sub>D src\<^sub>D \<mu>)
= \<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D \<eta> \<mu> \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D \<eta> (D.cod \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.cod \<mu>]) \<cdot>\<^sub>D
(I\<^sub>D.map \<mu> \<star>\<^sub>D src\<^sub>D \<mu>)
= \<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D (\<eta> (D.cod \<mu>) \<cdot>\<^sub>D \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
using \<mu> D.runit_naturality D.comp_assoc by simp
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D \<eta> \<mu> \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
using \<mu> \<eta>_naturality(1) by simp
finally show ?thesis by blast
qed
also have "... = (trg\<^sub>D \<mu> \<star>\<^sub>D FG.map \<mu>) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.dom \<mu>)] \<cdot>\<^sub>D \<eta> (D.dom \<mu>) \<cdot>\<^sub>D
\<r>\<^sub>D[D.dom \<mu>]"
proof -
have "\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D \<eta> \<mu> \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>] =
\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D (FG.map \<mu> \<cdot>\<^sub>D \<eta> (D.dom \<mu>)) \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
using \<mu> \<eta>_naturality(2) D.comp_assoc by simp
also have "... = (\<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \<mu>)] \<cdot>\<^sub>D FG.map \<mu>) \<cdot>\<^sub>D \<eta> (D.dom \<mu>) \<cdot>\<^sub>D
\<r>\<^sub>D[D.dom \<mu>]"
using D.comp_assoc by simp
also have "... = ((trg\<^sub>D \<mu> \<star>\<^sub>D FG.map \<mu>) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.dom \<mu>)]) \<cdot>\<^sub>D
\<eta> (D.dom \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
using \<mu> D.lunit'_naturality [of "FG.map \<mu>"] G\<^sub>0_props by simp
also have "... = (trg\<^sub>D \<mu> \<star>\<^sub>D FG.map \<mu>) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[FG.map (D.dom \<mu>)] \<cdot>\<^sub>D
\<eta> (D.dom \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.dom \<mu>]"
using D.comp_assoc by simp
finally show ?thesis by blast
qed
finally show "(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (D.cod \<mu>))] \<cdot>\<^sub>D \<eta> (D.cod \<mu>) \<cdot>\<^sub>D \<r>\<^sub>D[D.cod \<mu>]) \<cdot>\<^sub>D
(I\<^sub>D.map \<mu> \<star>\<^sub>D src\<^sub>D \<mu>)
= (trg\<^sub>D \<mu> \<star>\<^sub>D FG.map \<mu>) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G (D.dom \<mu>))] \<cdot>\<^sub>D \<eta> (D.dom \<mu>) \<cdot>\<^sub>D
\<r>\<^sub>D[D.dom \<mu>]"
by simp
qed
show "\<And>f g. \<lbrakk>D.ide f; D.ide g; src\<^sub>D g = trg\<^sub>D f\<rbrakk>
\<Longrightarrow> (trg\<^sub>D g \<star>\<^sub>D FG.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[trg\<^sub>D g, FG.map g, FG.map f] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G g)] \<cdot>\<^sub>D \<eta> g \<cdot>\<^sub>D \<r>\<^sub>D[g] \<star>\<^sub>D FG.map f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[I\<^sub>D.map g, src\<^sub>D g, FG.map f] \<cdot>\<^sub>D
(I\<^sub>D.map g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F (G f)] \<cdot>\<^sub>D \<eta> f \<cdot>\<^sub>D \<r>\<^sub>D[f]) \<cdot>\<^sub>D
\<a>\<^sub>D[I\<^sub>D.map g, I\<^sub>D.map f, src\<^sub>D f]
= (\<l>\<^sub>D\<^sup>-\<^sup>1[F (G (g \<star>\<^sub>D f))] \<cdot>\<^sub>D \<eta> (g \<star>\<^sub>D f) \<cdot>\<^sub>D
\<r>\<^sub>D[g \<star>\<^sub>D f]) \<cdot>\<^sub>D
(I\<^sub>D.cmp (g, f) \<star>\<^sub>D src\<^sub>D f)"
using \<eta>_respects_comp by simp
show "\<And>a. D.obj a \<Longrightarrow> (a \<star>\<^sub>D FG.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[a] \<cdot>\<^sub>D \<l>\<^sub>D[a] =
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (G a)] \<cdot>\<^sub>D \<eta> a \<cdot>\<^sub>D \<r>\<^sub>D[a]) \<cdot>\<^sub>D (I\<^sub>D.unit a \<star>\<^sub>D a)"
using \<eta>_respects_unit by auto
show "\<And>a. D.obj a \<Longrightarrow> D.equivalence_map a"
using D.obj_is_equivalence_map by simp
qed
abbreviation unit\<^sub>0'
where "unit\<^sub>0' \<equiv> \<lambda>a. a"
abbreviation unit\<^sub>1'
where "unit\<^sub>1' \<equiv> \<lambda>f. \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]"
interpretation \<eta>: pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
I\<^sub>C.map I\<^sub>C.cmp GF.map GF.cmp unit\<^sub>0' unit\<^sub>1'
proof
show "\<And>a. C.obj a \<Longrightarrow> C.ide a"
by auto
show "\<And>f. C.ide f \<Longrightarrow> C.iso (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)])"
using C.iso_runit C.iso_lunit'
by (intro C.isos_compose) auto
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>a : src\<^sub>C (I\<^sub>C.map a) \<rightarrow>\<^sub>C src\<^sub>C (GF.map a)\<guillemotright>"
by (simp_all add: C.obj_simps(1-3) G\<^sub>0_props(2))
show "\<And>f. C.ide f \<Longrightarrow> \<guillemotleft>\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)] :
GF.map f \<star>\<^sub>C src\<^sub>C f \<Rightarrow>\<^sub>C trg\<^sub>C f \<star>\<^sub>C I\<^sub>C.map f\<guillemotright>"
using G\<^sub>0_props by auto
show "\<And>\<mu>. C.arr \<mu> \<Longrightarrow>
(\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (C.cod \<mu>))]) \<cdot>\<^sub>C
(GF.map \<mu> \<star>\<^sub>C src\<^sub>C \<mu>)
= (trg\<^sub>C \<mu> \<star>\<^sub>C I\<^sub>C.map \<mu>) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[C.dom \<mu>] \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) \<cdot>\<^sub>C
\<r>\<^sub>C[G (F (C.dom \<mu>))]"
proof -
fix \<mu>
assume \<mu>: "C.arr \<mu>"
have "(\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.cod \<mu>)]) \<cdot>\<^sub>C
(GF.map \<mu> \<star>\<^sub>C src\<^sub>C \<mu>)
= \<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> \<mu> \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom \<mu>)]"
proof -
have "(\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.cod \<mu>)]) \<cdot>\<^sub>C
(GF.map \<mu> \<star>\<^sub>C src\<^sub>C \<mu>)
= \<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.cod \<mu>)] \<cdot>\<^sub>C
(GF.map \<mu> \<star>\<^sub>C src\<^sub>C \<mu>)"
using C.comp_assoc by simp
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C (\<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C GF.map \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom \<mu>)]"
using \<mu> C.runit_naturality [of "GF.map \<mu>"] G\<^sub>0_props C.comp_assoc by simp
also have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> \<mu> \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom \<mu>)]"
using \<mu> \<epsilon>_naturality(1) [of \<mu>] by simp
finally show ?thesis by blast
qed
also have "... = (trg\<^sub>C \<mu> \<star>\<^sub>C I\<^sub>C.map \<mu>) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[C.dom \<mu>] \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) \<cdot>\<^sub>C
\<r>\<^sub>C[GF.map (C.dom \<mu>)]"
proof -
have "... = \<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C (\<mu> \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>)) \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom \<mu>)]"
using \<mu> \<epsilon>_naturality(2) [of \<mu>] by simp
also have "... = (\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<mu>) \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[GF.map (C.dom \<mu>)]"
using C.comp_assoc by simp
also have "... = ((trg\<^sub>C \<mu> \<star>\<^sub>C I\<^sub>C.map \<mu>) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[C.dom \<mu>]) \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) \<cdot>\<^sub>C
\<r>\<^sub>C[GF.map (C.dom \<mu>)]"
using \<mu> C.lunit'_naturality [of \<mu>] by simp
also have "... = (trg\<^sub>C \<mu> \<star>\<^sub>C I\<^sub>C.map \<mu>) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[C.dom \<mu>] \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) \<cdot>\<^sub>C
\<r>\<^sub>C[GF.map (C.dom \<mu>)]"
using C.comp_assoc by simp
finally show ?thesis by blast
qed
finally show "(\<l>\<^sub>C\<^sup>-\<^sup>1[C.cod \<mu>] \<cdot>\<^sub>C \<epsilon> (C.cod \<mu>) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (C.cod \<mu>))]) \<cdot>\<^sub>C
(GF.map \<mu> \<star>\<^sub>C src\<^sub>C \<mu>)
= (trg\<^sub>C \<mu> \<star>\<^sub>C I\<^sub>C.map \<mu>) \<cdot>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[C.dom \<mu>] \<cdot>\<^sub>C \<epsilon> (C.dom \<mu>) \<cdot>\<^sub>C
\<r>\<^sub>C[G (F (C.dom \<mu>))]"
by simp
qed
show "\<And>f g. \<lbrakk>C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\<rbrakk> \<Longrightarrow>
(trg\<^sub>C g \<star>\<^sub>C I\<^sub>C.cmp (g, f)) \<cdot>\<^sub>C
\<a>\<^sub>C[trg\<^sub>C g, I\<^sub>C.map g, I\<^sub>C.map f] \<cdot>\<^sub>C
(\<l>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<epsilon> g \<cdot>\<^sub>C \<r>\<^sub>C[G (F g)] \<star>\<^sub>C I\<^sub>C.map f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[GF.map g, src\<^sub>C g, I\<^sub>C.map f] \<cdot>\<^sub>C
(GF.map g \<star>\<^sub>C \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<epsilon> f \<cdot>\<^sub>C \<r>\<^sub>C[G (F f)]) \<cdot>\<^sub>C
\<a>\<^sub>C[GF.map g, GF.map f, src\<^sub>C f]
= (\<l>\<^sub>C\<^sup>-\<^sup>1[g \<star>\<^sub>C f] \<cdot>\<^sub>C \<epsilon> (g \<star>\<^sub>C f) \<cdot>\<^sub>C \<r>\<^sub>C[G (F (g \<star>\<^sub>C f))]) \<cdot>\<^sub>C
(GF.cmp (g, f) \<star>\<^sub>C src\<^sub>C f)"
using \<epsilon>_respects_comp by simp
show "\<And>a. C.obj a \<Longrightarrow> (a \<star>\<^sub>C I\<^sub>C.unit a) \<cdot>\<^sub>C \<r>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<l>\<^sub>C[a] =
(\<l>\<^sub>C\<^sup>-\<^sup>1[a] \<cdot>\<^sub>C \<epsilon> a \<cdot>\<^sub>C \<r>\<^sub>C[G (F a)]) \<cdot>\<^sub>C (GF.unit a \<star>\<^sub>C a)"
using \<epsilon>_respects_unit by auto
show "\<And>a. C.obj a \<Longrightarrow> C.equivalence_map a"
using C.obj_is_equivalence_map by simp
qed
interpretation EQ: equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi> G \<Phi>\<^sub>G.map unit\<^sub>0' unit\<^sub>1' counit\<^sub>0' counit\<^sub>1'
..
lemma extends_to_equivalence_of_bicategories:
shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi> G \<Phi>\<^sub>G.map unit\<^sub>0' unit\<^sub>1' counit\<^sub>0' counit\<^sub>1'"
..
end
subsection "Equivalence Pseudofunctors Extend to Equivalences of Bicategories"
text \<open>
Now we put the pieces together and prove that an arbitrary equivalence pseudofunctor extends
to an equivalence of bicategories.
\<close>
context equivalence_pseudofunctor
begin
text \<open>
Define a set of objects \<open>U\<close> of \<open>C\<close> by choosing a representative of each equivalence
class of objects having the same image under the object map of the given equivalence
pseudofunctor. Then \<open>U\<close> is obviously dense, because every object of \<open>C\<close> belongs to
such an equivalence class.
\<close>
definition U
where "U = {a. C.obj a \<and> a = (SOME a'. C.obj a' \<and> map\<^sub>0 a' = map\<^sub>0 a)}"
lemma U_dense:
assumes "C.obj a"
shows "\<exists>a' \<in> U. C.equivalent_objects a a'"
proof -
let ?a' = "SOME a'. C.obj a' \<and> map\<^sub>0 a' = map\<^sub>0 a"
have "\<exists>a'. C.obj a' \<and> map\<^sub>0 a' = map\<^sub>0 a"
using assms by auto
hence 1: "?a' \<in> U \<and> map\<^sub>0 ?a' = map\<^sub>0 a"
using assms U_def someI_ex [of "\<lambda>a'. C.obj a' \<and> map\<^sub>0 a' = map\<^sub>0 a"] by simp
hence "C.equivalent_objects ?a' a"
using D.equivalent_objects_reflexive reflects_equivalent_objects [of ?a' a]
by (simp add: U_def assms)
thus ?thesis
using 1 C.equivalent_objects_symmetric by auto
qed
text \<open>
Take \<open>V\<close> to be the collection of images of all objects of \<open>C\<close> under the given equivalence
pseudofunctor. Since equivalence pseudofunctors are biessentially surjective on objects,
\<open>V\<close> is dense. Moreover, by construction, the object map of the given equivalence
pseudofunctor is a bijection from \<open>U\<close> to \<open>V\<close>.
\<close>
definition V
where "V = map\<^sub>0 ` Collect C.obj"
lemma V_dense:
assumes "D.obj b"
shows "\<exists>b'. b' \<in> map\<^sub>0 ` Collect C.obj \<and> D.equivalent_objects b b'"
using assms biessentially_surjective_on_objects D.equivalent_objects_symmetric
by blast
lemma bij_betw_U_V:
shows "bij_betw map\<^sub>0 U V"
proof -
have "inj_on map\<^sub>0 U"
using U_def by (intro inj_onI) simp
moreover have "map\<^sub>0 ` U = V"
proof
show "map\<^sub>0 ` U \<subseteq> V"
using U_def V_def by auto
show "V \<subseteq> map\<^sub>0 ` U"
proof
fix b
assume b: "b \<in> V"
obtain a where a: "C.obj a \<and> map\<^sub>0 a = b"
using b V_def by auto
let ?a' = "SOME a'. C.obj a' \<and> map\<^sub>0 a' = map\<^sub>0 a"
have 1: "C.obj ?a' \<and> map\<^sub>0 ?a' = b"
using a someI_ex [of "\<lambda>a'. C.obj a' \<and> map\<^sub>0 a' = map\<^sub>0 a"] by auto
moreover have "?a' = (SOME a''. C.obj a'' \<and> map\<^sub>0 a'' = map\<^sub>0 ?a')"
using a 1 by simp
ultimately have "?a' \<in> U"
unfolding U_def by simp
thus "b \<in> map\<^sub>0 ` U"
using a 1 U_def
by (metis (no_types, lifting) image_eqI)
qed
qed
ultimately show ?thesis
using bij_betw_def [of map\<^sub>0 U V] by simp
qed
abbreviation (input) Arr\<^sub>U
where "Arr\<^sub>U \<equiv> \<lambda>\<mu>. C.arr \<mu> \<and> src\<^sub>C \<mu> \<in> U \<and> trg\<^sub>C \<mu> \<in> U"
interpretation C\<^sub>U: subbicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C Arr\<^sub>U
using C.\<ll>_ide_simp C.\<rr>_ide_simp
apply unfold_locales
apply auto
apply (metis C.comp_ide_self C.ide_src C.src_cod C.src_dom)
by (metis C.trg.as_nat_trans.is_natural_2 C.trg.as_nat_trans.naturality C.trg_cod)
interpretation C\<^sub>U: dense_subbicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C U (* 15 sec *)
proof
show "\<And>a. C.obj a \<Longrightarrow> \<exists>a'. a' \<in> U \<and> C.equivalent_objects a' a"
using U_dense C.equivalent_objects_symmetric by blast
(* TODO: The above inference is trivial, but qed consumes 15 seconds! *)
qed (* 25 sec *)
abbreviation (input) Arr\<^sub>V
where "Arr\<^sub>V \<equiv> \<lambda>\<mu>. D.arr \<mu> \<and> src\<^sub>D \<mu> \<in> V \<and> trg\<^sub>D \<mu> \<in> V"
interpretation D\<^sub>V: subbicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D Arr\<^sub>V
using D.\<ll>_ide_simp D.\<rr>_ide_simp
apply unfold_locales
apply auto
apply (metis D.comp_ide_self D.ide_src D.src_cod D.src_dom)
by (metis D.trg.as_nat_trans.is_natural_2 D.trg.as_nat_trans.naturality D.trg_cod)
interpretation D\<^sub>V: dense_subbicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V (* 25 sec *)
using V_dense D.equivalent_objects_def D.equivalent_objects_symmetric V_def
by unfold_locales metis (* 35 sec -- time is similar doing it this way. *)
interpretation F\<^sub>U: restricted_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>
\<open>\<lambda>\<mu>. C.arr \<mu> \<and> src\<^sub>C \<mu> \<in> U \<and> trg\<^sub>C \<mu> \<in> U\<close>
..
interpretation F\<^sub>U\<^sub>V: corestricted_pseudofunctor
C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F\<^sub>U.map F\<^sub>U.cmp \<open>\<lambda>a. a \<in> V\<close>
proof
show "\<And>\<mu>. C\<^sub>U.arr \<mu> \<Longrightarrow> D\<^sub>V.arr (F\<^sub>U.map \<mu>)"
proof -
fix \<mu>
assume \<mu>: "C\<^sub>U.arr \<mu>"
have "src\<^sub>C \<mu> \<in> U \<and> trg\<^sub>C \<mu> \<in> U"
- using \<mu> C\<^sub>U.arr_char by simp
+ using \<mu> C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "\<And>a. a \<in> U \<Longrightarrow> F\<^sub>U.map\<^sub>0 a = map\<^sub>0 a"
proof -
fix a :: 'a
assume a: "a \<in> U"
hence 1: "C.obj a"
using U_def by blast
have "src\<^sub>C a = a"
using a U_def by blast
thus "F\<^sub>U.map\<^sub>0 a = map\<^sub>0 a"
- using a 1 C.obj_simps(1,3) C\<^sub>U.arr_char F\<^sub>U.map\<^sub>0_def map\<^sub>0_def by presburger
+ using a 1 C.obj_simps(1,3) C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C F\<^sub>U.map\<^sub>0_def map\<^sub>0_def by presburger
qed
ultimately have "F\<^sub>U.map\<^sub>0 (src\<^sub>C \<mu>) \<in> V \<and> F\<^sub>U.map\<^sub>0 (trg\<^sub>C \<mu>) \<in> V"
using \<mu> bij_betw_U_V bij_betw_def
by (metis (no_types, lifting) image_eqI)
hence "src\<^sub>D (F\<^sub>U.map \<mu>) \<in> V \<and> trg\<^sub>D (F\<^sub>U.map \<mu>) \<in> V"
using \<mu> F\<^sub>U.map\<^sub>0_def C\<^sub>U.src_def C\<^sub>U.trg_def F\<^sub>U.preserves_src F\<^sub>U.preserves_trg
by auto
thus "D\<^sub>V.arr (F\<^sub>U.map \<mu>)"
- using \<mu> D\<^sub>V.arr_char [of "F\<^sub>U.map \<mu>"] by blast
+ using \<mu> D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C [of "F\<^sub>U.map \<mu>"] by blast
qed
qed
interpretation F\<^sub>U\<^sub>V: equivalence_pseudofunctor
C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg
F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp
proof
show "\<And>f f'. \<lbrakk>C\<^sub>U.par f f'; F\<^sub>U.map f = F\<^sub>U.map f'\<rbrakk> \<Longrightarrow> f = f'"
- using C\<^sub>U.arr_char C\<^sub>U.dom_char C\<^sub>U.cod_char
+ using C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.dom_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) is_faithful)
show "\<And>a b g. \<lbrakk>C\<^sub>U.obj a; C\<^sub>U.obj b; D\<^sub>V.in_hhom g (F\<^sub>U\<^sub>V.map\<^sub>0 a) (F\<^sub>U\<^sub>V.map\<^sub>0 b);
D\<^sub>V.ide g\<rbrakk>
\<Longrightarrow> \<exists>f. C\<^sub>U.in_hhom f a b \<and> C\<^sub>U.ide f \<and> D\<^sub>V.isomorphic (F\<^sub>U.map f) g"
proof -
fix a b g
assume a: "C\<^sub>U.obj a" and b: "C\<^sub>U.obj b"
assume g: "D\<^sub>V.in_hhom g (F\<^sub>U\<^sub>V.map\<^sub>0 a) (F\<^sub>U\<^sub>V.map\<^sub>0 b)" and ide_g: "D\<^sub>V.ide g"
have 1: "\<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) g"
proof -
have "C.obj a \<and> C.obj b"
using a b C\<^sub>U.obj_char by simp
moreover have "\<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 b\<guillemotright>"
- using a b g D\<^sub>V.in_hhom_def D\<^sub>V.arr_char D\<^sub>V.src_def D\<^sub>V.trg_def
+ using a b g D\<^sub>V.in_hhom_def D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.src_def D\<^sub>V.trg_def
by (intro D.in_hhomI) auto
moreover have "D.ide g"
- using ide_g D\<^sub>V.ide_char D\<^sub>V.arr_char by simp
+ using ide_g D\<^sub>V.ide_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
ultimately show ?thesis
using locally_essentially_surjective by simp
qed
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) g"
using 1 by blast
have 2: "C\<^sub>U.in_hhom f a b"
- using f a b C\<^sub>U.arr_char C\<^sub>U.obj_char C\<^sub>U.src_def C\<^sub>U.trg_def by fastforce
+ using f a b C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.obj_char C\<^sub>U.src_def C\<^sub>U.trg_def by fastforce
moreover have "C\<^sub>U.ide f"
- using f 2 C\<^sub>U.ide_char C\<^sub>U.arr_char by auto
+ using f 2 C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
moreover have "D\<^sub>V.isomorphic (F\<^sub>U\<^sub>V.map f) g"
proof -
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : F f \<Rightarrow>\<^sub>D g\<guillemotright> \<and> D.iso \<phi>"
using f D.isomorphic_def by auto
have 3: "D\<^sub>V.in_hom \<phi> (F\<^sub>U\<^sub>V.map f) g"
proof -
have "D\<^sub>V.arr \<phi>"
- using f g \<phi> 2 D\<^sub>V.arr_char
+ using f g \<phi> 2 D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) D.arrI D.vconn_implies_hpar(1-4) D\<^sub>V.ideD(1) ide_g)
thus ?thesis
- using \<phi> 2 D\<^sub>V.dom_char D\<^sub>V.cod_char
+ using \<phi> 2 D\<^sub>V.dom_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.cod_char\<^sub>S\<^sub>b\<^sub>C
by (intro D\<^sub>V.in_homI) auto
qed
moreover have "D\<^sub>V.iso \<phi>"
- using \<phi> D\<^sub>V.iso_char D\<^sub>V.arr_char 3 by fastforce
+ using \<phi> D\<^sub>V.iso_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C 3 by fastforce
ultimately show ?thesis
using D\<^sub>V.isomorphic_def by auto
qed
ultimately show "\<exists>f. C\<^sub>U.in_hhom f a b \<and> C\<^sub>U.ide f \<and> D\<^sub>V.isomorphic (F\<^sub>U\<^sub>V.map f) g"
by auto
qed
show "\<And>f f' \<nu>. \<lbrakk>C\<^sub>U.ide f; C\<^sub>U.ide f'; C\<^sub>U.src f = C\<^sub>U.src f'; C\<^sub>U.trg f = C\<^sub>U.trg f';
D\<^sub>V.in_hom \<nu> (F\<^sub>U.map f) (F\<^sub>U.map f')\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. C\<^sub>U.in_hom \<mu> f f' \<and> F\<^sub>U.map \<mu> = \<nu>"
proof -
fix f f' \<nu>
assume f: "C\<^sub>U.ide f" and f': "C\<^sub>U.ide f'"
and eq_src: "C\<^sub>U.src f = C\<^sub>U.src f'" and eq_trg: "C\<^sub>U.trg f = C\<^sub>U.trg f'"
and \<nu>: "D\<^sub>V.in_hom \<nu> (F\<^sub>U.map f) (F\<^sub>U.map f')"
have 1: "C.ide f \<and> C.ide f'"
- using f f' C\<^sub>U.ide_char C\<^sub>U.arr_char by simp
+ using f f' C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have 2: "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> F \<mu> = \<nu>"
proof -
have "src\<^sub>C f = src\<^sub>C f' \<and> trg\<^sub>C f = trg\<^sub>C f'"
using f f' 1 eq_src eq_trg C\<^sub>U.src_def C\<^sub>U.trg_def by simp
moreover have "\<guillemotleft>\<nu> : F f \<Rightarrow>\<^sub>D F f'\<guillemotright>"
- using \<nu> D\<^sub>V.in_hom_char D\<^sub>V.arr_char by auto
+ using \<nu> D\<^sub>V.in_hom_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show ?thesis
using 1 locally_full by simp
qed
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> F \<mu> = \<nu>"
using 2 by auto
have 3: "C\<^sub>U.in_hom \<mu> f f'"
- using \<mu> f f' C\<^sub>U.in_hom_char C\<^sub>U.ide_char C\<^sub>U.arr_char by auto
+ using \<mu> f f' C\<^sub>U.in_hom_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
moreover have "F\<^sub>U.map \<mu> = \<nu>"
using \<mu> 3 by auto
ultimately show "\<exists>\<mu>. C\<^sub>U.in_hom \<mu> f f' \<and> F\<^sub>U.map \<mu> = \<nu>"
by auto
qed
show "\<And>b. D\<^sub>V.obj b \<Longrightarrow> \<exists>a. C\<^sub>U.obj a \<and> D\<^sub>V.equivalent_objects (F\<^sub>U\<^sub>V.map\<^sub>0 a) b"
proof -
fix b
assume b: "D\<^sub>V.obj b"
obtain a where a: "C.obj a \<and> map\<^sub>0 a = b"
- using b D\<^sub>V.obj_char D\<^sub>V.arr_char V_def by auto
+ using b D\<^sub>V.obj_char D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C V_def by auto
have 1: "D.obj b \<and> b \<in> V"
using a b D\<^sub>V.obj_char V_def by auto
obtain a' where a': "a' \<in> U \<and> C.equivalent_objects a' a"
using a U_dense C.equivalent_objects_symmetric by blast
have obj_a': "C\<^sub>U.obj a'"
- using a' U_def C\<^sub>U.obj_char C\<^sub>U.arr_char by fastforce
+ using a' U_def C\<^sub>U.obj_char C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
moreover have "D\<^sub>V.equivalent_objects (F\<^sub>U\<^sub>V.map\<^sub>0 a') b"
proof -
have "D.equivalent_objects (map\<^sub>0 a') (map\<^sub>0 a)"
using a' preserves_equivalent_objects by simp
hence 2: "D.equivalent_objects (map\<^sub>0 a') b"
using a D.equivalent_objects_transitive by simp
obtain e where e: "\<guillemotleft>e : map\<^sub>0 a' \<rightarrow>\<^sub>D b\<guillemotright> \<and> D.equivalence_map e"
using 2 D.equivalent_objects_def by auto
have 3: "D.obj (map\<^sub>0 a') \<and> map\<^sub>0 a' \<in> V"
using a' e U_def V_def by simp
moreover have e_in_hhom: "D\<^sub>V.in_hhom e (F\<^sub>U\<^sub>V.map\<^sub>0 a') b"
proof
show 4: "D\<^sub>V.arr e"
- using 1 3 a e b D\<^sub>V.arr_char D\<^sub>V.obj_char V_def
+ using 1 3 a e b D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.obj_char V_def
by (metis (no_types, lifting) D.in_hhomE)
show "D\<^sub>V.src e = F\<^sub>U\<^sub>V.map\<^sub>0 a'"
using e 4 D\<^sub>V.src_def F\<^sub>U\<^sub>V.map\<^sub>0_def obj_a' map\<^sub>0_def by auto
show "D\<^sub>V.trg e = b"
using e 4 D\<^sub>V.trg_def by auto
qed
moreover have "D\<^sub>V.equivalence_map e"
proof -
obtain d \<eta> \<epsilon> where d: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D e d \<eta> \<epsilon>"
using e D.equivalence_map_def by auto
interpret e: equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D e d \<eta> \<epsilon>
using d by simp
have "equivalence_in_bicategory D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg e d \<eta> \<epsilon>"
proof
show ide_e: "D\<^sub>V.ide e"
- using e e_in_hhom D\<^sub>V.ide_char by auto
+ using e e_in_hhom D\<^sub>V.ide_char\<^sub>S\<^sub>b\<^sub>C by auto
show ide_d: "D\<^sub>V.ide d"
- using d ide_e e.antipar D\<^sub>V.ide_char D\<^sub>V.arr_char by simp
+ using d ide_e e.antipar D\<^sub>V.ide_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show 4: "D\<^sub>V.in_hom \<eta> (D\<^sub>V.src e) (D\<^sub>V.hcomp d e)"
proof -
have "D\<^sub>V.hseq d e"
using ide_d ide_e e.antipar D\<^sub>V.src_def D\<^sub>V.trg_def by simp
thus ?thesis
- using ide_d ide_e e.unit_in_hom D\<^sub>V.arr_char D\<^sub>V.ide_char
- D\<^sub>V.dom_char D\<^sub>V.cod_char D\<^sub>V.src_def D\<^sub>V.trg_def
+ using ide_d ide_e e.unit_in_hom D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.ide_char\<^sub>S\<^sub>b\<^sub>C
+ D\<^sub>V.dom_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.cod_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.src_def D\<^sub>V.trg_def
by (intro D\<^sub>V.in_homI) auto
qed
show 5: "D\<^sub>V.in_hom \<epsilon> (D\<^sub>V.hcomp e d) (D\<^sub>V.src d)"
proof -
have "D\<^sub>V.hseq e d"
using ide_d ide_e e.antipar D\<^sub>V.src_def D\<^sub>V.trg_def by simp
thus ?thesis
- using ide_d ide_e e.antipar e.counit_in_hom D\<^sub>V.arr_char D\<^sub>V.ide_char
- D\<^sub>V.dom_char D\<^sub>V.cod_char D\<^sub>V.src_def D\<^sub>V.trg_def
+ using ide_d ide_e e.antipar e.counit_in_hom D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.ide_char\<^sub>S\<^sub>b\<^sub>C
+ D\<^sub>V.dom_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.cod_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.src_def D\<^sub>V.trg_def
by (intro D\<^sub>V.in_homI) auto
qed
show "D\<^sub>V.iso \<eta>"
- using 4 D\<^sub>V.iso_char D\<^sub>V.arr_char by fastforce
+ using 4 D\<^sub>V.iso_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
show "D\<^sub>V.iso \<epsilon>"
- using 5 D\<^sub>V.iso_char D\<^sub>V.arr_char by fastforce
+ using 5 D\<^sub>V.iso_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
qed
thus ?thesis
using D\<^sub>V.equivalence_map_def by auto
qed
ultimately show ?thesis
using D\<^sub>V.equivalent_objects_def by auto
qed
ultimately show "\<exists>a. C\<^sub>U.obj a \<and> D\<^sub>V.equivalent_objects (F\<^sub>U\<^sub>V.map\<^sub>0 a) b" by auto
qed
qed
interpretation F\<^sub>U\<^sub>V: equivalence_pseudofunctor_bij_on_obj
C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg
F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp
proof
have "Collect C\<^sub>U.obj = U"
- using C\<^sub>U.obj_char C\<^sub>U.arr_char U_def by fastforce
+ using C\<^sub>U.obj_char C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C U_def by fastforce
moreover have "Collect D\<^sub>V.obj = V"
- using D\<^sub>V.obj_char D\<^sub>V.arr_char V_def D.obj_def' map\<^sub>0_simps(1) by auto
+ using D\<^sub>V.obj_char D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C V_def D.obj_def' map\<^sub>0_simps(1) by auto
moreover have "\<And>a. a \<in> Collect C\<^sub>U.obj \<Longrightarrow> F\<^sub>U\<^sub>V.map\<^sub>0 a = map\<^sub>0 a"
- using F\<^sub>U\<^sub>V.map\<^sub>0_def C\<^sub>U.obj_char C\<^sub>U.arr_char D\<^sub>V.src_def
+ using F\<^sub>U\<^sub>V.map\<^sub>0_def C\<^sub>U.obj_char C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.src_def
F\<^sub>U.map\<^sub>0_simp F\<^sub>U\<^sub>V.map\<^sub>0_simp by auto
ultimately show "bij_betw F\<^sub>U\<^sub>V.map\<^sub>0 (Collect C\<^sub>U.obj) (Collect D\<^sub>V.obj)"
using bij_betw_U_V
by (simp add: bij_betw_U_V bij_betw_cong)
qed
interpretation EQ\<^sub>U\<^sub>V: equivalence_of_bicategories (* 25 sec *)
D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg
C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp F\<^sub>U\<^sub>V.G F\<^sub>U\<^sub>V.\<Phi>\<^sub>G
F\<^sub>U\<^sub>V.unit\<^sub>0' F\<^sub>U\<^sub>V.unit\<^sub>1' F\<^sub>U\<^sub>V.counit\<^sub>0' F\<^sub>U\<^sub>V.counit\<^sub>1'
using F\<^sub>U\<^sub>V.extends_to_equivalence_of_bicategories by blast (* 18 sec, mostly "by" *)
text \<open>
Now compose \<open>EQ\<^sub>U\<^sub>V\<close> with the equivalence from \<open>D\<^sub>V\<close> to \<open>D\<close> and the converse of the equivalence
from \<open>C\<^sub>U\<close> to \<open>C\<close>. The result is an equivalence of bicategories from \<open>C\<close> to \<open>D\<close>.
\<close>
interpretation EQ\<^sub>C: equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
C\<^sub>U.E C\<^sub>U.\<Phi>\<^sub>E C\<^sub>U.P C\<^sub>U.\<Phi>\<^sub>P
C\<^sub>U.unit\<^sub>0 C\<^sub>U.unit\<^sub>1 C\<^sub>U.counit\<^sub>0 C\<^sub>U.counit\<^sub>1
using C\<^sub>U.induces_equivalence by simp
interpretation EQ\<^sub>C': converse_equivalence_of_bicategories
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
C\<^sub>U.E C\<^sub>U.\<Phi>\<^sub>E C\<^sub>U.P C\<^sub>U.\<Phi>\<^sub>P
C\<^sub>U.unit\<^sub>0 C\<^sub>U.unit\<^sub>1 C\<^sub>U.counit\<^sub>0 C\<^sub>U.counit\<^sub>1
..
interpretation EQ\<^sub>D: equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg
D\<^sub>V.E D\<^sub>V.\<Phi>\<^sub>E D\<^sub>V.P D\<^sub>V.\<Phi>\<^sub>P
D\<^sub>V.unit\<^sub>0 D\<^sub>V.unit\<^sub>1 D\<^sub>V.counit\<^sub>0 D\<^sub>V.counit\<^sub>1
using D\<^sub>V.induces_equivalence by simp
interpretation EQ\<^sub>U\<^sub>VoEQ\<^sub>C': composite_equivalence_of_bicategories (* 35 sec *)
D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg
C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\<a> \<i>\<^sub>C C\<^sub>U.src C\<^sub>U.trg
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp F\<^sub>U\<^sub>V.G F\<^sub>U\<^sub>V.\<Phi>\<^sub>G
C\<^sub>U.P C\<^sub>U.\<Phi>\<^sub>P C\<^sub>U.E C\<^sub>U.\<Phi>\<^sub>E
F\<^sub>U\<^sub>V.unit\<^sub>0' F\<^sub>U\<^sub>V.unit\<^sub>1' F\<^sub>U\<^sub>V.counit\<^sub>0' F\<^sub>U\<^sub>V.counit\<^sub>1'
EQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>C'.counit\<^sub>1
.. (* 55 sec *)
interpretation EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C': composite_equivalence_of_bicategories (* 30 sec *)
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\<a> \<i>\<^sub>D D\<^sub>V.src D\<^sub>V.trg
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
D\<^sub>V.E D\<^sub>V.\<Phi>\<^sub>E D\<^sub>V.P D\<^sub>V.\<Phi>\<^sub>P
EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp
EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp
D\<^sub>V.unit\<^sub>0 D\<^sub>V.unit\<^sub>1 D\<^sub>V.counit\<^sub>0 D\<^sub>V.counit\<^sub>1
EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1
EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1
.. (* 50 sec *)
lemma induces_equivalence_of_bicategories:
shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1"
..
lemma left_map_simp:
assumes "C.arr \<mu>"
shows "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \<mu> = D\<^sub>V.E (F (C\<^sub>U.P \<mu>))"
using assms by simp
lemma right_map_simp:
assumes "D.arr \<nu>"
shows "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map \<nu> = C\<^sub>U.E (F\<^sub>U\<^sub>V.G (D\<^sub>V.P \<nu>))"
using assms by simp
lemma unit\<^sub>0_simp:
assumes "C.obj a"
shows "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 a =
C\<^sub>U.E (F\<^sub>U\<^sub>V.G (D\<^sub>V.\<eta>\<^sub>0 (D\<^sub>V.src (F (C\<^sub>U.P a))))) \<star>\<^sub>C C\<^sub>U.E (C\<^sub>U.P\<^sub>0 (src\<^sub>C a))
\<star>\<^sub>C EQ\<^sub>C'.unit\<^sub>0 a"
using assms EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0_simp EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0_simp
EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_def C\<^sub>U.src_def
by auto
(*
* TODO: Expansions of the other parts of the equivalence are not comprehensible
* without interpreting a bunch of other locales. I don't know that there is any particular
* value to doing that. The above show that the pseudofunctors between C and D and the
* components of the unit of the equivalence are as expected.
*)
text \<open>
We've now got an equivalence of bicategories between \<open>C\<close> and \<open>D\<close>, but it involves
\<open>EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map\<close> and not the originally given equivalence pseudofunctor \<open>F\<close>.
However, we can patch things up by showing that \<open>EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map\<close> is pseudonaturally
equivalent to \<open>F\<close>. From this, we may conclude, using the fact that equivalences of
bicategories respect pseudonatural equivalence, that there is an equivalence of bicategories
between \<open>C\<close> and \<open>D\<close> that involves \<open>F\<close> and \<open>EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map\<close>, rather than
\<open>EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map\<close> and \<open>EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map\<close>.
\<close>
abbreviation \<tau>\<^sub>0
where "\<tau>\<^sub>0 a \<equiv> F (C\<^sub>U.\<epsilon>\<^sub>0 a)"
abbreviation \<tau>\<^sub>1
where "\<tau>\<^sub>1 f \<equiv> D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)) \<cdot>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 f) \<cdot>\<^sub>D \<Phi> (f, C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C f))"
lemma \<tau>\<^sub>0_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>\<tau>\<^sub>0 a : map\<^sub>0 (C\<^sub>U.P\<^sub>0 a) \<rightarrow>\<^sub>D map\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<tau>\<^sub>0 a : \<tau>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0 a\<guillemotright>"
proof -
show "\<guillemotleft>\<tau>\<^sub>0 a : map\<^sub>0 (C\<^sub>U.P\<^sub>0 a) \<rightarrow>\<^sub>D map\<^sub>0 a\<guillemotright>"
proof
show "D.arr (\<tau>\<^sub>0 a)"
using assms by simp
show "src\<^sub>D (\<tau>\<^sub>0 a) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)"
using assms map\<^sub>0_def C\<^sub>U.counit.ide_map\<^sub>0_obj C\<^sub>U.equivalence_data_simps\<^sub>B(7) C.ideD(1)
preserves_src
by presburger
show "trg\<^sub>D (\<tau>\<^sub>0 a) = map\<^sub>0 a"
using assms by auto
qed
show "\<guillemotleft>\<tau>\<^sub>0 a : \<tau>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0 a\<guillemotright>"
using assms by auto
qed
lemma \<tau>\<^sub>0_simps [simp]:
assumes "C.obj a"
shows "D.ide (\<tau>\<^sub>0 a)"
and "src\<^sub>D (\<tau>\<^sub>0 a) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" and "trg\<^sub>D (\<tau>\<^sub>0 a) = map\<^sub>0 a"
using assms \<tau>\<^sub>0_in_hom(1) [of a] by auto
lemma \<tau>\<^sub>1_in_hom [intro]:
assumes "C.ide f"
shows "\<guillemotleft>\<tau>\<^sub>1 f : map\<^sub>0 (C\<^sub>U.P\<^sub>0 (src\<^sub>C f)) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C f)\<guillemotright>"
and "\<guillemotleft>\<tau>\<^sub>1 f : F f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F (C\<^sub>U.P f)\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<tau>\<^sub>1 f : F f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F (C\<^sub>U.P f)\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<Phi> (f, C\<^sub>U.d (src\<^sub>C f)) : F f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D F (f \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C f))\<guillemotright>"
using assms by auto
show "\<guillemotleft>F (C\<^sub>U.\<epsilon>\<^sub>1 f) : F (f \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C f)) \<Rightarrow>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>C C\<^sub>U.P f)\<guillemotright>"
using assms C\<^sub>U.counit\<^sub>1_in_hom [of f] C\<^sub>U.P_def by auto
show "\<guillemotleft>D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)) :
F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>C C\<^sub>U.P f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F (C\<^sub>U.P f)\<guillemotright>"
- using assms C.VV.arr_char C.VV.ide_char \<Phi>.components_are_iso
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C \<Phi>.components_are_iso
by (metis (no_types, lifting) C\<^sub>U.P_def C\<^sub>U.counit.ide_map\<^sub>0_obj C\<^sub>U.counit\<^sub>1_simps(1,5)
C\<^sub>U.equivalence_data_simps\<^sub>B(2) C.hseqE C.ide_hcomp C.obj_trg D.arr_cod D.inv_in_homI
cmp_components_are_iso cmp_in_hom(2) preserves_cod preserves_reflects_arr)
qed
show "\<guillemotleft>\<tau>\<^sub>1 f : map\<^sub>0 (C\<^sub>U.P\<^sub>0 (src\<^sub>C f)) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C f)\<guillemotright>"
using assms 1 map\<^sub>0_def [of "C\<^sub>U.P\<^sub>0 (src\<^sub>C f)"] C\<^sub>U.emb.map\<^sub>0_simp C\<^sub>U.P\<^sub>0_props
apply (intro D.in_hhomI)
apply auto
using D.src_dom [of "\<tau>\<^sub>1 f"]
apply (elim D.in_homE)
apply auto
using D.trg_dom [of "\<tau>\<^sub>1 f"]
apply (elim D.in_homE)
by auto
qed
lemma \<tau>\<^sub>1_simps [simp]:
assumes "C.ide f"
shows "D.arr (\<tau>\<^sub>1 f)"
and "src\<^sub>D (\<tau>\<^sub>1 f) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 (src\<^sub>C f))" and "trg\<^sub>D (\<tau>\<^sub>1 f) = map\<^sub>0 (trg\<^sub>C f)"
and "D.dom (\<tau>\<^sub>1 f) = F f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)" and "D.cod (\<tau>\<^sub>1 f) = \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F (C\<^sub>U.P f)"
using assms \<tau>\<^sub>1_in_hom by blast+
lemma iso_\<tau>\<^sub>1:
assumes "C.ide f"
shows "D.iso (\<tau>\<^sub>1 f)"
proof -
have "C.VV.ide (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"
proof -
have "C.VV.arr (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"
using assms C\<^sub>U.equivalence_data_simps\<^sub>B(7) C.ideD(1) by auto
moreover have "C.VxV.ide (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"
- using assms C\<^sub>U.prj.preserves_ide [of f] C\<^sub>U.ide_char by simp
+ using assms C\<^sub>U.prj.preserves_ide [of f] C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C by simp
ultimately show ?thesis
- using assms C.VV.ide_char [of "(C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"] by blast
+ using assms C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C [of "(C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"] by blast
qed
hence "D.iso (\<Phi> (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))"
by simp
thus ?thesis
- using assms C.VV.ide_char C.VV.arr_char \<Phi>.components_are_iso
+ using assms C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>.components_are_iso
by (intro D.isos_compose) auto
qed
interpretation \<tau>: pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp F \<Phi>
\<tau>\<^sub>0 \<tau>\<^sub>1
proof
show "\<And>a. C.obj a \<Longrightarrow> D.ide (\<tau>\<^sub>0 a)"
by simp
show "\<And>a. C.obj a \<Longrightarrow> D.equivalence_map (\<tau>\<^sub>0 a)"
using C\<^sub>U.counit.components_are_equivalences preserves_equivalence_maps by blast
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>\<tau>\<^sub>0 a : src\<^sub>D (EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map a) \<rightarrow>\<^sub>D src\<^sub>D (F a)\<guillemotright>"
proof -
fix a
assume a: "C.obj a"
show "\<guillemotleft>\<tau>\<^sub>0 a : src\<^sub>D (EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map a) \<rightarrow>\<^sub>D src\<^sub>D (F a)\<guillemotright>"
proof
show "D.arr (\<tau>\<^sub>0 a)"
using a by simp
show "src\<^sub>D (\<tau>\<^sub>0 a) = src\<^sub>D (EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map a)"
proof -
have "D\<^sub>V.arr (F (C\<^sub>U.P a))"
proof -
have "src\<^sub>D (F (C\<^sub>U.P a)) \<in> V"
proof -
have "src\<^sub>D (F (C\<^sub>U.P a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
using a by auto
moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)"
using a C\<^sub>U.obj_char [of "C\<^sub>U.prj.map\<^sub>0 a"] C\<^sub>U.prj.map\<^sub>0_simps(1) by auto
ultimately show ?thesis
using V_def by simp
qed
moreover have "trg\<^sub>D (F (C\<^sub>U.P a)) \<in> V"
proof -
have "trg\<^sub>D (F (C\<^sub>U.P a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
using a by auto
moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)"
using a C\<^sub>U.obj_char [of "C\<^sub>U.prj.map\<^sub>0 a"] C\<^sub>U.prj.map\<^sub>0_simps(1) by auto
ultimately show ?thesis
using V_def by simp
qed
ultimately show ?thesis
- using a D\<^sub>V.arr_char by fastforce
+ using a D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
qed
thus ?thesis
using a D\<^sub>V.emb.map\<^sub>0_def D\<^sub>V.emb.map_def D\<^sub>V.src_def C.obj_simps(1-2) C\<^sub>U.P_simps\<^sub>B(2)
by (simp add: C\<^sub>U.P\<^sub>0_props(1))
qed
show "trg\<^sub>D (F (C\<^sub>U.d a)) = src\<^sub>D (F a)"
using a by auto
qed
qed
show "\<And>f. C.ide f \<Longrightarrow>
\<guillemotleft>\<tau>\<^sub>1 f : F f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f\<guillemotright>"
proof -
fix f
assume f: "C.ide f"
show "\<guillemotleft>\<tau>\<^sub>1 f : F f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f\<guillemotright>"
proof -
have "D\<^sub>V.arr (F (C\<^sub>U.P f))"
using f F\<^sub>U\<^sub>V.preserves_arr by auto
hence "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f = F (C\<^sub>U.P f)"
using f EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_def D\<^sub>V.emb.map_def by simp
thus ?thesis
using f by auto
qed
qed
show "\<And>f. C.ide f \<Longrightarrow> D.iso (\<tau>\<^sub>1 f)"
using iso_\<tau>\<^sub>1 by simp
show "\<And>\<mu>. C.arr \<mu> \<Longrightarrow>
\<tau>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) =
(\<tau>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 (C.dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "C.arr \<mu>"
show "\<tau>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>))
= (\<tau>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 (C.dom \<mu>)"
proof -
have "\<tau>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>))
= D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P (C.cod \<mu>))) \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D
\<Phi> (C.cod \<mu>, C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
(F \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>))"
using \<mu> D.comp_assoc by simp
also have "... = D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P (C.cod \<mu>))) \<cdot>\<^sub>D
(F (C\<^sub>U.\<epsilon>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D
F (\<mu> \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))) \<cdot>\<^sub>D
\<Phi> (C.dom \<mu>, C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))"
using \<mu> \<Phi>.naturality [of "(\<mu>, C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))"] D.comp_assoc
- C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
also have "... = (D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P (C.cod \<mu>))) \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>C C\<^sub>U.EoP.map \<mu>)) \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>1 (C.dom \<mu>)) \<cdot>\<^sub>D
\<Phi> (C.dom \<mu>, C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))"
proof -
have "F (C\<^sub>U.\<epsilon>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D F (\<mu> \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))
= F (C\<^sub>U.\<epsilon>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>C (\<mu> \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>)))"
using \<mu> by simp
also have "... = F ((C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>C C\<^sub>U.EoP.map \<mu>) \<cdot>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>1 (C.dom \<mu>))"
using \<mu> C\<^sub>U.counit.naturality [of \<mu>] by simp
also have "... = F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>C C\<^sub>U.EoP.map \<mu>) \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>1 (C.dom \<mu>))"
using \<mu> by simp
finally have "F (C\<^sub>U.\<epsilon>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D F (\<mu> \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))
= F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>C C\<^sub>U.EoP.map \<mu>) \<cdot>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 (C.dom \<mu>))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>)) \<star>\<^sub>D F (C\<^sub>U.EoP.map \<mu>)) \<cdot>\<^sub>D
D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P (C.dom \<mu>))) \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>1 (C.dom \<mu>)) \<cdot>\<^sub>D
\<Phi> (C.dom \<mu>, C\<^sub>U.\<epsilon>\<^sub>0 (src\<^sub>C \<mu>))"
proof -
have "D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P (C.cod \<mu>))) \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>C C\<^sub>U.EoP.map \<mu>)
= (F (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>)) \<star>\<^sub>D F (C\<^sub>U.EoP.map \<mu>)) \<cdot>\<^sub>D
D.inv (\<Phi> (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P (C.dom \<mu>)))"
proof -
have "C.VV.arr (C\<^sub>U.\<epsilon>\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P \<mu>)"
- proof (unfold C.VV.arr_char, intro conjI)
+ proof (unfold C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C, intro conjI)
show "C.arr (fst (C\<^sub>U.d (trg\<^sub>C \<mu>), C\<^sub>U.P \<mu>))"
using \<mu> by simp
show "C.arr (snd (C\<^sub>U.d (trg\<^sub>C \<mu>), C\<^sub>U.P \<mu>))"
using \<mu> by simp
show "src\<^sub>C (fst (C\<^sub>U.d (trg\<^sub>C \<mu>), C\<^sub>U.P \<mu>)) = trg\<^sub>C (snd (C\<^sub>U.d (trg\<^sub>C \<mu>), C\<^sub>U.P \<mu>))"
using \<mu> C\<^sub>U.emb.map\<^sub>0_def C\<^sub>U.emb.map_def apply simp
using C\<^sub>U.P\<^sub>0_props(1) C\<^sub>U.P_simps\<^sub>B(3) C\<^sub>U.src_def by fastforce
qed
moreover have "C\<^sub>U.EoP.map \<mu> = C\<^sub>U.P \<mu>"
using \<mu> by (simp add: C\<^sub>U.emb.map_def)
moreover have "C\<^sub>U.emb.map\<^sub>0 (C\<^sub>U.P\<^sub>0 (trg\<^sub>C \<mu>)) = C\<^sub>U.P\<^sub>0 (trg\<^sub>C \<mu>)"
using \<mu> C\<^sub>U.P\<^sub>0_props(1) C\<^sub>U.emb.map\<^sub>0_simp by blast
ultimately show ?thesis
- using \<mu> C.VV.arr_char C.VV.dom_char C.VV.cod_char C.VV.ide_char FF_def
+ using \<mu> C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
\<Phi>.inv_naturality [of "(C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \<mu>), C\<^sub>U.P \<mu>)"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 (C.dom \<mu>)"
using \<mu> F\<^sub>U\<^sub>V.preserves_arr D.comp_assoc D\<^sub>V.emb.map_def C\<^sub>U.emb.map_def by simp
finally show "\<tau>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>))
= (\<tau>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 (C.dom \<mu>)"
by blast
qed
qed
show "\<And>a. C.obj a \<Longrightarrow>
(\<tau>\<^sub>0 a \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a] =
\<tau>\<^sub>1 a \<cdot>\<^sub>D (unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
fix a
assume a: "C.obj a"
have 1: "C\<^sub>U.obj (C\<^sub>U.prj.map\<^sub>0 a)"
using a C\<^sub>U.prj.map\<^sub>0_simps(1) by simp
have 2: "\<guillemotleft>C\<^sub>U.prj.unit a : C\<^sub>U.prj.map\<^sub>0 a \<Rightarrow>\<^sub>C C\<^sub>U.P a\<guillemotright>"
- using a C\<^sub>U.in_hom_char by blast
+ using a C\<^sub>U.in_hom_char\<^sub>S\<^sub>b\<^sub>C by blast
have 3: "\<guillemotleft>C\<^sub>U.prj.unit a : C\<^sub>U.prj.map\<^sub>0 a \<rightarrow>\<^sub>C C\<^sub>U.prj.map\<^sub>0 a\<guillemotright>"
proof (intro C.in_hhomI)
show 4: "C.arr (C\<^sub>U.prj.unit a)"
using 2 by auto
show "src\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a"
proof -
have "src\<^sub>C (C\<^sub>U.prj.unit a) = src\<^sub>C (C\<^sub>U.P a)"
using 2 4 C.src_cod [of "C\<^sub>U.prj.unit a"] C.vconn_implies_hpar(1,3) by auto
also have "... = C\<^sub>U.prj.map\<^sub>0 a"
using a C.obj_simps(1) C\<^sub>U.prj.map\<^sub>0_def [of a] C\<^sub>U.src_def [of "C\<^sub>U.P a"]
by simp
finally show ?thesis by blast
qed
show "trg\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a"
proof -
have "trg\<^sub>C (C\<^sub>U.prj.unit a) = trg\<^sub>C (C\<^sub>U.P a)"
using 2 4 C.trg_cod [of "C\<^sub>U.prj.unit a"] C.vconn_implies_hpar(2,4) by auto
also have "... = C\<^sub>U.prj.map\<^sub>0 a"
using a C.obj_simps(1,3) C\<^sub>U.prj.map\<^sub>0_def [of a] C\<^sub>U.trg_def [of "C\<^sub>U.P a"]
by simp
finally show ?thesis by blast
qed
qed
have [simp]: "C.dom (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a"
using 2 by auto
have [simp]: "C.arr (C\<^sub>U.prj.unit a)"
using 2 by auto
have [simp]: "C.cod (C\<^sub>U.prj.unit a) = C\<^sub>U.P a"
using 2 by auto
have [simp]: "src\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a"
using 3 by auto
have [simp]: "trg\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a"
using 3 by auto
have [simp]: "C.arr (C\<^sub>U.d a)"
using a by simp
have [simp]: "C.ide (C\<^sub>U.d a)"
using a by simp
have [simp]: "C.dom (C\<^sub>U.d a) = C\<^sub>U.d a"
using a by auto
have [simp]: "C.cod (C\<^sub>U.d a) = C\<^sub>U.d a"
using a by auto
have [simp]: "src\<^sub>C (C\<^sub>U.d a) = C\<^sub>U.prj.map\<^sub>0 a"
using a C\<^sub>U.equivalence_data_simps\<^sub>B(7) by auto
have [simp]: "trg\<^sub>C (C\<^sub>U.d a) = a"
using a C\<^sub>U.equivalence_data_simps\<^sub>B(7) by auto
have seq: "D.seq (F (C\<^sub>U.prj.unit a)) (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
using a
apply (intro D.seqI)
- using C\<^sub>U.prj.map\<^sub>0_simps(1) D\<^sub>V.arr_char F\<^sub>U\<^sub>V.unit_simps(1) apply blast
- using C\<^sub>U.prj.unit_simps(1) C\<^sub>U.arr_char apply blast
+ using C\<^sub>U.prj.map\<^sub>0_simps(1) D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C F\<^sub>U\<^sub>V.unit_simps(1) apply blast
+ using C\<^sub>U.prj.unit_simps(1) C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C apply blast
proof -
have "D.dom (F (C\<^sub>U.prj.unit a)) = F (C.dom (C\<^sub>U.prj.unit a))"
- using a C\<^sub>U.prj.unit_simps(1) C\<^sub>U.arr_char by simp
+ using a C\<^sub>U.prj.unit_simps(1) C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = F (C\<^sub>U.prj.map\<^sub>0 a)"
- using a C\<^sub>U.prj.unit_simps [of a] C\<^sub>U.dom_char by simp
+ using a C\<^sub>U.prj.unit_simps [of a] C\<^sub>U.dom_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = D.cod (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
- using a 1 C\<^sub>U.prj.unit_simps [of a] F\<^sub>U\<^sub>V.unit_simps(5) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.cod_char
+ using a 1 C\<^sub>U.prj.unit_simps [of a] F\<^sub>U\<^sub>V.unit_simps(5) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally show "D.dom (F (C\<^sub>U.prj.unit a)) = D.cod (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
by blast
qed
have 4: "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a = F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)"
proof -
have "EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)"
using a EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_def D\<^sub>V.src_def C.obj_def F\<^sub>U\<^sub>V.preserves_arr by auto
have seq': "D\<^sub>V.arr (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
proof -
have 5: "D.arr (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
using seq by simp
moreover have "src\<^sub>D (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) \<in> V"
proof -
have "src\<^sub>D (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))
= map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
proof -
have "src\<^sub>D (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))
= src\<^sub>D (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
using 5 D.src_vcomp D.vseq_implies_hpar by simp
also have "... = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
using a F\<^sub>U\<^sub>V.unit_simps(1-2) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.src_def F\<^sub>U\<^sub>V.map\<^sub>0_def
apply simp
using a C\<^sub>U.prj.map\<^sub>0_simps(1) F\<^sub>U\<^sub>V.preserves_arr map\<^sub>0_def by force
finally show ?thesis by blast
qed
moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)"
using a C\<^sub>U.prj.map\<^sub>0_simps(1) C\<^sub>U.obj_char by blast
ultimately show ?thesis
using V_def by blast
qed
moreover have "trg\<^sub>D (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) \<in> V"
proof -
have "trg\<^sub>D (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))
= map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
proof -
have "trg\<^sub>D (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) =
trg\<^sub>D (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
using 5 D.trg_vcomp D.vseq_implies_hpar by simp
also have "... = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
using a F\<^sub>U\<^sub>V.unit_simps(1,3) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.src_def D\<^sub>V.trg_def
F\<^sub>U\<^sub>V.map\<^sub>0_def
apply simp
using a C\<^sub>U.prj.map\<^sub>0_simps(1) F\<^sub>U\<^sub>V.preserves_arr map\<^sub>0_def by force
finally show ?thesis by blast
qed
moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)"
using a C\<^sub>U.prj.map\<^sub>0_simps(1) C\<^sub>U.obj_char by blast
ultimately show ?thesis
using V_def by blast
qed
ultimately show ?thesis
- using a D\<^sub>V.arr_char by simp
+ using a D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a =
F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a) \<cdot>\<^sub>D map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)"
proof -
have 5: "D\<^sub>V.obj (src\<^sub>D (F (C\<^sub>U.P\<^sub>0 a)))"
using a C\<^sub>U.P\<^sub>0_props [of a]
by (metis (no_types, lifting) EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_simps(1)
\<open>EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)\<close> map\<^sub>0_def)
have 6: "D\<^sub>V.emb.unit (map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)"
using a 5 D\<^sub>V.emb.unit_char' EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_simps(1)
\<open>EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)\<close> map\<^sub>0_def
by simp
moreover have "D\<^sub>V.emb.unit (map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)"
using 6 a C.obj_simps C\<^sub>U.equivalence_data_simps\<^sub>B(7) by simp
moreover have "D\<^sub>V.arr (F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
using a C.obj_simps seq' by blast
moreover have "D\<^sub>V.arr (F\<^sub>U\<^sub>V.unit (C\<^sub>U.P\<^sub>0 a))"
using a 1 by simp
moreover have "D\<^sub>V.arr (F (C\<^sub>U.prj.unit a))"
using a F\<^sub>U\<^sub>V.preserves_arr by auto
moreover have "D\<^sub>V.obj (F\<^sub>U\<^sub>V.map\<^sub>0 (C\<^sub>U.P\<^sub>0 a))"
using a 1 F\<^sub>U\<^sub>V.map\<^sub>0_simps(1) by auto
moreover have "F\<^sub>U\<^sub>V.map\<^sub>0 (C\<^sub>U.P\<^sub>0 a) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)"
using a \<open>EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)\<close> map\<^sub>0_def F\<^sub>U\<^sub>V.map\<^sub>0_simps
by auto
ultimately show ?thesis
using a seq EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit_char' [of a] EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit_char' [of a]
- D\<^sub>V.emb.map_def D\<^sub>V.seq_char D\<^sub>V.comp_char D.comp_assoc
+ D\<^sub>V.emb.map_def D\<^sub>V.seq_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.comp_char D.comp_assoc
by simp
qed
also have "... = F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)"
proof -
have "D.dom (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)"
- using a 1 F\<^sub>U\<^sub>V.unit_simps(4) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.dom_char F\<^sub>U\<^sub>V.map\<^sub>0_def
+ using a 1 F\<^sub>U\<^sub>V.unit_simps(4) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.dom_char\<^sub>S\<^sub>b\<^sub>C F\<^sub>U\<^sub>V.map\<^sub>0_def
D\<^sub>V.src_def map\<^sub>0_def C.obj_def F\<^sub>U\<^sub>V.preserves_arr
by simp
moreover have "D.arr (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))"
using a \<open>D.seq (F (C\<^sub>U.prj.unit a)) (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))\<close> by blast
ultimately show ?thesis
using a D.comp_arr_dom by auto
qed
also have "... = F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)"
using a 1 F\<^sub>U\<^sub>V.unit_char' F\<^sub>U.unit_char' by simp
finally show ?thesis by blast
qed
have "(\<tau>\<^sub>0 a \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a] =
(\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a]"
using 4 by simp
also have "(\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a]
= (\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D
unit (C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 a \<star>\<^sub>D D.inv (unit (src\<^sub>C (C\<^sub>U.d a)))) \<cdot>\<^sub>D
D.inv (\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D
F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>D
F \<l>\<^sub>C[C\<^sub>U.d a] \<cdot>\<^sub>D
\<Phi> (trg\<^sub>C (C\<^sub>U.d a), C\<^sub>U.d a) \<cdot>\<^sub>D
(unit (trg\<^sub>C (C\<^sub>U.d a)) \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "\<l>\<^sub>D[\<tau>\<^sub>0 a] = F \<l>\<^sub>C[C\<^sub>U.d a] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C (C\<^sub>U.d a), C\<^sub>U.d a) \<cdot>\<^sub>D
(unit (trg\<^sub>C (C\<^sub>U.d a)) \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a preserves_lunit [of "C\<^sub>U.\<epsilon>\<^sub>0 a"] C\<^sub>U.counit.ide_map\<^sub>0_obj lunit_coherence
by blast
moreover
have "\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] = (\<tau>\<^sub>0 a \<star>\<^sub>D D.inv (unit (src\<^sub>C (C\<^sub>U.d a)))) \<cdot>\<^sub>D
D.inv (\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a]"
proof -
have "F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] =
\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a)) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]"
using preserves_runit(2) [of "C\<^sub>U.\<epsilon>\<^sub>0 a"] by simp
moreover have 1: "D.iso (\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a)))"
- using a C.VV.ide_char C.VV.arr_char C.obj_src [of "C\<^sub>U.d a"]
+ using a C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.obj_src [of "C\<^sub>U.d a"]
\<Phi>.components_are_iso [of "(C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))"]
by auto
ultimately have "D.inv (\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] =
(\<tau>\<^sub>0 a \<star>\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]"
using D.invert_side_of_triangle(1)
[of "F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a]" "\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))"
"(\<tau>\<^sub>0 a \<star>\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]"]
by fastforce
thus ?thesis
using D.invert_side_of_triangle(1)
[of "D.inv (\<Phi> (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \<cdot>\<^sub>D F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a]"
"\<tau>\<^sub>0 a \<star>\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))" "\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]"]
using C.obj_src [of "C\<^sub>U.d a"] unit_char(2) by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0 a \<star>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 a \<star>\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a))))) \<cdot>\<^sub>D
D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D
F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>D F \<l>\<^sub>C[C\<^sub>U.d a] \<cdot>\<^sub>D
\<Phi> (a, C\<^sub>U.d a) \<cdot>\<^sub>D
(unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a) =
(\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a))"
using 1 seq D.whisker_left [of "\<tau>\<^sub>0 a"]
by (simp add: F\<^sub>U.unit_char' F\<^sub>U\<^sub>V.unit_char')
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D
D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.prj.map\<^sub>0 a))) \<cdot>\<^sub>D
F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>D F \<l>\<^sub>C[C\<^sub>U.d a] \<cdot>\<^sub>D
\<Phi> (a, C\<^sub>U.d a) \<cdot>\<^sub>D
(unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "(\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0 a \<star>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 a \<star>\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a))))
= (\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 a \<star>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a) \<cdot>\<^sub>D
D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a)))"
proof -
have "D.seq (unit (C\<^sub>U.prj.map\<^sub>0 a)) (D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a)))"
using a 1 unit_char(1-2) C\<^sub>U.obj_char
by (intro D.seqI) auto
thus ?thesis
using a D.whisker_left [of "\<tau>\<^sub>0 a"] by simp
qed
also have "... = (\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.map\<^sub>0 a))"
using a 1 unit_char(1-2) C\<^sub>U.obj_char unit_simps [of "C\<^sub>U.prj.map\<^sub>0 a"]
by (simp add: D.comp_arr_inv')
also have "... = \<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F (C\<^sub>U.prj.map\<^sub>0 a)"
using seq D.whisker_left by auto
also have "... = \<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)"
using D.comp_arr_dom by simp
finally have "(\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0 a \<star>\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 a \<star>\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a))))
= \<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.P a)) \<cdot>\<^sub>D
(F (C\<^sub>U.d a \<star>\<^sub>C C\<^sub>U.prj.unit a) \<cdot>\<^sub>D
F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>D F \<l>\<^sub>C[C\<^sub>U.d a]) \<cdot>\<^sub>D
\<Phi> (a, C\<^sub>U.d a) \<cdot>\<^sub>D
(unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "(\<tau>\<^sub>0 a \<star>\<^sub>D F (C\<^sub>U.prj.unit a)) \<cdot>\<^sub>D D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.prj.map\<^sub>0 a)) =
D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.P a)) \<cdot>\<^sub>D F (C\<^sub>U.d a \<star>\<^sub>C C\<^sub>U.prj.unit a)"
proof -
have "C.VV.arr (C\<^sub>U.d a, C\<^sub>U.prj.unit a)"
- using a C.VV.arr_char by simp
+ using a C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
thus ?thesis
using a \<Phi>.inv_naturality [of "(C\<^sub>U.d a, C\<^sub>U.prj.unit a)"] C\<^sub>U.prj.unit_simps [of a]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.P a)) \<cdot>\<^sub>D
F (C\<^sub>U.counit\<^sub>1 a \<cdot>\<^sub>C (a \<star>\<^sub>C C\<^sub>U.d a)) \<cdot>\<^sub>D
\<Phi> (a, C\<^sub>U.d a)) \<cdot>\<^sub>D
(unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "F (C\<^sub>U.d a \<star>\<^sub>C C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>D F \<l>\<^sub>C[C\<^sub>U.d a] =
F ((C\<^sub>U.d a \<star>\<^sub>C C\<^sub>U.prj.unit a) \<cdot>\<^sub>C \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>C \<l>\<^sub>C[C\<^sub>U.d a])"
using a by simp
also have "... = F (C\<^sub>U.counit\<^sub>1 a \<cdot>\<^sub>C (a \<star>\<^sub>C C\<^sub>U.d a))"
proof -
have "C\<^sub>U.EoP.unit a = C\<^sub>U.prj.unit a"
using a 1 C\<^sub>U.emb.map_def C\<^sub>U.EoP.unit_char' C\<^sub>U.emb.unit_char' C.comp_arr_dom
by simp
thus ?thesis
using a C\<^sub>U.counit.respects_unit [of a] C\<^sub>U.I\<^sub>C.unit_char' [of a] by simp
qed
finally have "F (C\<^sub>U.d a \<star>\<^sub>C C\<^sub>U.prj.unit a) \<cdot>\<^sub>D F \<r>\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \<cdot>\<^sub>D F \<l>\<^sub>C[C\<^sub>U.d a] =
F (C\<^sub>U.counit\<^sub>1 a \<cdot>\<^sub>C (a \<star>\<^sub>C C\<^sub>U.d a))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<tau>\<^sub>1 a \<cdot>\<^sub>D (unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.P a)) \<cdot>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 a \<cdot>\<^sub>C (a \<star>\<^sub>C C\<^sub>U.d a)) \<cdot>\<^sub>D \<Phi> (a, C\<^sub>U.d a)
= D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.P a)) \<cdot>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 a) \<cdot>\<^sub>D \<Phi> (a, C\<^sub>U.d a)"
using a C.obj_def' C.comp_arr_dom
by (metis C\<^sub>U.counit\<^sub>1_simps(1) C\<^sub>U.counit\<^sub>1_simps(4) C.objE)
also have "... = \<tau>\<^sub>1 a"
using a by auto
finally have "D.inv (\<Phi> (C\<^sub>U.d a, C\<^sub>U.P a)) \<cdot>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 a \<cdot>\<^sub>C (a \<star>\<^sub>C C\<^sub>U.d a)) \<cdot>\<^sub>D \<Phi> (a, C\<^sub>U.d a)
= \<tau>\<^sub>1 a"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
finally show "(\<tau>\<^sub>0 a \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a]
= \<tau>\<^sub>1 a \<cdot>\<^sub>D (unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
by blast
qed
show "\<And>f g. \<lbrakk>C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\<rbrakk> \<Longrightarrow>
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g,
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (src\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]
= \<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))"
proof -
fix f g
assume f: "C.ide f" and g: "C.ide g" and fg: "src\<^sub>C g = trg\<^sub>C f"
have 1: "C.ide f \<and> C.ide g \<and> C.ide (C\<^sub>U.P f) \<and> C.ide (C\<^sub>U.P g) \<and>
C.ide (C\<^sub>U.d (trg\<^sub>C f)) \<and> C.ide (C\<^sub>U.d (trg\<^sub>C g)) \<and>
src\<^sub>C (C\<^sub>U.d (trg\<^sub>C f)) = trg\<^sub>C (C\<^sub>U.P f) \<and>
src\<^sub>C (C\<^sub>U.d (trg\<^sub>C g)) = trg\<^sub>C (C\<^sub>U.P g) \<and>
trg\<^sub>C (C\<^sub>U.d (trg\<^sub>C f)) = src\<^sub>C g"
using f g fg C\<^sub>U.emb.map\<^sub>0_def C\<^sub>U.emb.map_def C\<^sub>U.obj_char C\<^sub>U.P\<^sub>0_props(1) C.obj_simps(2)
- C\<^sub>U.prj.preserves_ide C\<^sub>U.ide_char
+ C\<^sub>U.prj.preserves_ide C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C
by auto
have "\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))
= (D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
F ((C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>C
\<a>\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \<cdot>\<^sub>C
(C\<^sub>U.\<epsilon>\<^sub>1 g \<star>\<^sub>C C\<^sub>U.P f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \<cdot>\<^sub>C
(g \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>1 f) \<cdot>\<^sub>C
\<a>\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)]) \<cdot>\<^sub>D
\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
(\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))"
using f g fg C\<^sub>U.emb.map_def C\<^sub>U.counit.respects_hcomp [of f g] D.comp_arr_dom
by simp
also have "... = ((D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
(\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))"
proof -
have "F ((C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>C
\<a>\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \<cdot>\<^sub>C
(C\<^sub>U.\<epsilon>\<^sub>1 g \<star>\<^sub>C C\<^sub>U.P f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \<cdot>\<^sub>C
(g \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>1 f) \<cdot>\<^sub>C
\<a>\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)])
= F (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \<cdot>\<^sub>D
F (C\<^sub>U.\<epsilon>\<^sub>1 g \<star>\<^sub>C C\<^sub>U.P f) \<cdot>\<^sub>D
F (C.inv \<a>\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f]) \<cdot>\<^sub>D
F (g \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>1 f) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)]"
proof -
have "C.arr ((C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>C
\<a>\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \<cdot>\<^sub>C
(C\<^sub>U.\<epsilon>\<^sub>1 g \<star>\<^sub>C C\<^sub>U.P f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \<cdot>\<^sub>C
(g \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>1 f) \<cdot>\<^sub>C
\<a>\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)])"
- using f g fg 1 C\<^sub>U.emb.map_def C.VV.arr_char C.VV.dom_char C\<^sub>U.EoP.FF_def
+ using f g fg 1 C\<^sub>U.emb.map_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
by (intro C.seqI C.hseqI') auto
thus ?thesis
using f g fg by fastforce
qed
also have "... = \<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \<cdot>\<^sub>D
((D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g)) \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
((D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
(F (C\<^sub>U.\<epsilon>\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f))) \<cdot>\<^sub>D
((D.inv (\<Phi> (g \<star>\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (g \<star>\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f)) \<cdot>\<^sub>D
(\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f)) \<star>\<^sub>D F (C\<^sub>U.P f))) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(F g \<star>\<^sub>D D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \<cdot>\<^sub>D
((D.inv (\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f) \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f) \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 f))) \<cdot>\<^sub>D
((D.inv (\<Phi> (g, f \<star>\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (g, f \<star>\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<Phi> (f, C\<^sub>U.d (src\<^sub>C f)))) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))"
using 1 f g fg preserves_hcomp preserves_assoc C\<^sub>U.emb.map_def
- C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def D.comp_assoc
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def D.comp_assoc
by simp
also have "... = \<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
((D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g)) \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
(F (C\<^sub>U.\<epsilon>\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
(\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f)) \<star>\<^sub>D F (C\<^sub>U.P f))) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
((F g \<star>\<^sub>D D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \<cdot>\<^sub>D
(F g \<star>\<^sub>D F (C\<^sub>U.\<epsilon>\<^sub>1 f)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<Phi> (f, C\<^sub>U.d (src\<^sub>C f)))) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))"
proof -
have "(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))
= \<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
- using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg 1 C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
moreover have "(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
(F (C\<^sub>U.\<epsilon>\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f))
= F (C\<^sub>U.\<epsilon>\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f)"
proof -
have "(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
(F (C\<^sub>U.\<epsilon>\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f))
= (F (C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.P g) \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
(F (C\<^sub>U.\<epsilon>\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f))"
- using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg 1 C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
also have "... = F (C\<^sub>U.counit\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f)"
using g D.comp_cod_arr f g fg C\<^sub>U.P\<^sub>0_props(1) C\<^sub>U.emb.map_def by simp
finally show ?thesis by blast
qed
moreover have "(D.inv (\<Phi> (g \<star>\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f)) \<cdot>\<^sub>D
(\<Phi> (g \<star>\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \<cdot>\<^sub>D
(\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f)) \<star>\<^sub>D F (C\<^sub>U.P f))
= (\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f)) \<star>\<^sub>D F (C\<^sub>U.P f))"
- using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg 1 C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
moreover have "(D.inv (\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f) \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f) \<star>\<^sub>C C\<^sub>U.P f)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D F (C\<^sub>U.counit\<^sub>1 f))
= F g \<star>\<^sub>D F (C\<^sub>U.counit\<^sub>1 f)"
- using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg 1 C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
moreover have "(D.inv (\<Phi> (g, f \<star>\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (g, f \<star>\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<Phi> (f, C\<^sub>U.d (src\<^sub>C f)))
= F g \<star>\<^sub>D \<Phi> (f, C\<^sub>U.d (src\<^sub>C f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))"
proof -
have "(F g \<star>\<^sub>D D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \<cdot>\<^sub>D
(F g \<star>\<^sub>D F (C\<^sub>U.counit\<^sub>1 f)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<Phi> (f, C\<^sub>U.d (src\<^sub>C f)))
= F g \<star>\<^sub>D \<tau>\<^sub>1 f"
- using f g fg 1 D.whisker_left C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg 1 D.whisker_left C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g)) \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
(F (C\<^sub>U.counit\<^sub>1 g) \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
(\<Phi> (g, C\<^sub>U.d (trg\<^sub>C f)) \<star>\<^sub>D F (C\<^sub>U.P f))
= \<tau>\<^sub>1 g \<star>\<^sub>D F (C\<^sub>U.P f)"
- using f g fg 1 D.whisker_right C.VV.arr_char C.VV.dom_char C.VV.cod_char
- C.VV.ide_char \<Phi>.components_are_iso C\<^sub>U.emb.map_def
+ using f g fg 1 D.whisker_right C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C \<Phi>.components_are_iso C\<^sub>U.emb.map_def
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
finally have "F ((C\<^sub>U.d (trg\<^sub>C g) \<star>\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>C
\<a>\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \<cdot>\<^sub>C
(C\<^sub>U.\<epsilon>\<^sub>1 g \<star>\<^sub>C C\<^sub>U.P f) \<cdot>\<^sub>C
C.inv \<a>\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \<cdot>\<^sub>C
(g \<star>\<^sub>C C\<^sub>U.\<epsilon>\<^sub>1 f) \<cdot>\<^sub>C
\<a>\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)])
= \<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]"
proof -
have "(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)))
= \<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))"
proof -
have "(D.inv (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)))
= (\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.P (g \<star>\<^sub>C f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)))"
proof -
have "D.iso (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f)))"
- using f g fg \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char
- by (metis (no_types, lifting) C\<^sub>U.prj.preserves_ide C\<^sub>U.P_simps\<^sub>B(3) C\<^sub>U.ide_char
+ using f g fg \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis (no_types, lifting) C\<^sub>U.prj.preserves_ide C\<^sub>U.P_simps\<^sub>B(3) C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C
C\<^sub>U.counit.ide_map\<^sub>0_obj C\<^sub>U.equivalence_data_simps\<^sub>B(7) C.hcomp_simps(2)
C.ideD(1) C.ide_hcomp C.obj_trg cmp_components_are_iso)
moreover have "D.dom (\<Phi> (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \<star>\<^sub>C f))) =
\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.P (g \<star>\<^sub>C f))"
- using f g fg C\<^sub>U.ide_char C\<^sub>U.equivalence_data_simps\<^sub>B(7) C\<^sub>U.prj.preserves_ide
+ using f g fg C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.equivalence_data_simps\<^sub>B(7) C\<^sub>U.prj.preserves_ide
cmp_simps(4)
by simp
ultimately show ?thesis
using D.comp_inv_arr' by auto
qed
also have "... = \<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.whisker_left [of "\<tau>\<^sub>0 (trg\<^sub>C g)"]
by simp
also have "... = \<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_cod_arr
by simp
finally show ?thesis by blast
qed
moreover have "\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D
(\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))
= \<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]"
proof -
have "D.inv (\<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)) =
F (g \<star>\<^sub>C f) \<star>\<^sub>D F (C\<^sub>U.d (src\<^sub>C f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
moreover have "(F (g \<star>\<^sub>C f) \<star>\<^sub>D F (C\<^sub>U.d (src\<^sub>C f))) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) =
\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
by simp
moreover have "(D.inv (\<Phi> (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) =
(F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
- \<Phi>.components_are_iso C.VV.ide_char FF_def
+ \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def
D.whisker_right [of "\<tau>\<^sub>0 (src\<^sub>C f)" "D.inv (\<Phi> (g, f))" "\<Phi> (g, f)"]
by simp
moreover have "\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D ((F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) =
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
by simp
ultimately show ?thesis by argo
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F (C\<^sub>U.P f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]"
proof -
have "(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))
= \<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
proof -
have "D.arr (F (C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.EoP.FF_def
C\<^sub>U.emb.map_def
by (intro D.seqI) auto
thus ?thesis
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
D.whisker_left [of "\<tau>\<^sub>0 (trg\<^sub>C g)"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g,
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (src\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]"
proof -
have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f) =
F (C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
proof -
have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f) =
F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D ((F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.\<Phi>\<^sub>P (g, f))) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D D\<^sub>V.\<Phi>\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f))"
proof -
have 3: "D.arr (F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def C\<^sub>U.prj.FF_def
- C\<^sub>U.\<Phi>\<^sub>P_in_hom [of g f] C\<^sub>U.hcomp_def C\<^sub>U.in_hom_char
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def C\<^sub>U.prj.FF_def
+ C\<^sub>U.\<Phi>\<^sub>P_in_hom [of g f] C\<^sub>U.hcomp_def C\<^sub>U.in_hom_char\<^sub>S\<^sub>b\<^sub>C
by (intro D.seqI) auto
have 4: "D\<^sub>V.arr (\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
using f g fg 1 cmp_in_hom(1) [of "C\<^sub>U.P g" "C\<^sub>U.P f"] C\<^sub>U.prj.map\<^sub>0_simps(1)
- C\<^sub>U.obj_char D\<^sub>V.arr_char V_def
+ C\<^sub>U.obj_char D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C V_def
by auto
moreover have 5: "D\<^sub>V.arr (F (C\<^sub>U.\<Phi>\<^sub>P (g, f)))"
using f g fg 1 C\<^sub>U.\<Phi>\<^sub>P_simps [of g f] F\<^sub>U\<^sub>V.preserves_arr by presburger
moreover have 6: "D\<^sub>V.arr (F (C\<^sub>U.P (g \<star>\<^sub>C f)))"
using f g fg C\<^sub>U.P_simps(1) [of "g \<star>\<^sub>C f"] F\<^sub>U\<^sub>V.preserves_arr by simp
moreover have 7: "D\<^sub>V.arr (F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
- using 3 4 5 D\<^sub>V.seq_char [of "F (C\<^sub>U.\<Phi>\<^sub>P (g, f))" "\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"]
+ using 3 4 5 D\<^sub>V.seq_char\<^sub>S\<^sub>b\<^sub>C [of "F (C\<^sub>U.\<Phi>\<^sub>P (g, f))" "\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"]
D\<^sub>V.comp_char [of "F (C\<^sub>U.\<Phi>\<^sub>P (g, f))" "\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"]
by auto
moreover have "D.seq (F (C\<^sub>U.\<Phi>\<^sub>P (g, f))) (\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
using 3 by blast
moreover have "D.seq (F (C\<^sub>U.P (g \<star>\<^sub>C f)))
(F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
using 3 by blast
moreover have "D\<^sub>V.arr (F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f))"
- using f g fg 3 6 7 D.vseq_implies_hpar D\<^sub>V.arr_char V_def
- by (metis (no_types, lifting) D\<^sub>V.comp_char D\<^sub>V.seq_char)
+ using f g fg 3 6 7 D.vseq_implies_hpar V_def
+ by (metis (no_types, lifting) D\<^sub>V.comp_char D\<^sub>V.seq_char\<^sub>S\<^sub>b\<^sub>C)
ultimately show ?thesis
using f g fg EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.cmp_def EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.cmp_def
- C.VV.arr_char C.VV.dom_char C\<^sub>U.VV.arr_char D\<^sub>V.comp_char
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.comp_char
D\<^sub>V.emb.map_def D.comp_assoc
by simp
qed
also have "... = F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f) \<cdot>\<^sub>D D\<^sub>V.\<Phi>\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f))"
proof -
have "F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) = F (C\<^sub>U.\<Phi>\<^sub>P (g, f))"
proof -
have 1: "C.arr (C\<^sub>U.\<Phi>\<^sub>P (g, f))"
- using f g fg C.VV.arr_char
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) C\<^sub>U.prj.preserves_hcomp C\<^sub>U.prj.preserves_ide
- C\<^sub>U.ide_char C\<^sub>U.seq_char C.ideD(1) C.ideD(3) C.ide_hcomp C.seqE)
+ C\<^sub>U.ide_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.seq_char\<^sub>S\<^sub>b\<^sub>C C.ideD(1) C.ideD(3) C.ide_hcomp C.seqE)
moreover have "C.cod (C\<^sub>U.\<Phi>\<^sub>P (g, f)) = C\<^sub>U.P (g \<star>\<^sub>C f)"
- using f g fg C.VV.arr_char C\<^sub>U.\<Phi>\<^sub>P_in_hom(2) [of g f]
- C\<^sub>U.hcomp_def C\<^sub>U.in_hom_char by auto
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.\<Phi>\<^sub>P_in_hom(2) [of g f]
+ C\<^sub>U.hcomp_def C\<^sub>U.in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately have "D.cod (F (C\<^sub>U.\<Phi>\<^sub>P (g, f))) = F (C\<^sub>U.P (g \<star>\<^sub>C f))"
- using f g fg C.VV.arr_char preserves_cod [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C preserves_cod [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
by simp
thus ?thesis
using 1 f g fg D.comp_cod_arr [of "F (C\<^sub>U.\<Phi>\<^sub>P (g, f))" "F (C\<^sub>U.P (g \<star>\<^sub>C f))"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = F (C\<^sub>U.P (g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (C\<^sub>U.\<Phi>\<^sub>P (g, f)) \<cdot>\<^sub>D
F (C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
proof -
have "D\<^sub>V.\<Phi>\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f)) = F (C\<^sub>U.P g) \<star>\<^sub>D F (C\<^sub>U.P f)"
- using f g fg D\<^sub>V.emb.cmp_def D\<^sub>V.VV.arr_char D\<^sub>V.src_def D\<^sub>V.trg_def
+ using f g fg D\<^sub>V.emb.cmp_def D\<^sub>V.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D\<^sub>V.src_def D\<^sub>V.trg_def
F\<^sub>U\<^sub>V.preserves_arr
by auto
moreover have "C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f) = C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f"
- using f g fg C\<^sub>U.VV.arr_char C\<^sub>U.emb.cmp_def by simp
+ using f g fg C\<^sub>U.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.emb.cmp_def by simp
ultimately
have "\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f) \<cdot>\<^sub>D D\<^sub>V.\<Phi>\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f)) =
F (C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
\<Phi>.naturality [of "(C\<^sub>U.P g, C\<^sub>U.P f)"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = F (C\<^sub>U.P (g \<star>\<^sub>C f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>P (g, f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f)) \<cdot>\<^sub>D
\<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
proof -
have "C.arr (C\<^sub>U.P (g \<star>\<^sub>C f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>P (g, f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
proof (intro C.seqI)
show "C.arr (C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
using f g fg by auto
show 1: "C.arr (C\<^sub>U.\<Phi>\<^sub>P (g, f))"
- using f g fg C\<^sub>U.arr_char C\<^sub>U.\<Phi>\<^sub>P_simps(1) [of g f] by auto
+ using f g fg C\<^sub>U.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.\<Phi>\<^sub>P_simps(1) [of g f] by auto
show "C.dom (C\<^sub>U.\<Phi>\<^sub>P (g, f)) = C.cod (C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
proof -
have "C.dom (C\<^sub>U.\<Phi>\<^sub>P (g, f)) = C\<^sub>U.hcomp (C\<^sub>U.P g) (C\<^sub>U.P f)"
- using f g fg C\<^sub>U.VV.arr_char C\<^sub>U.\<Phi>\<^sub>P_simps(4) [of g f]
- C\<^sub>U.dom_char [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
+ using f g fg C\<^sub>U.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.\<Phi>\<^sub>P_simps(4) [of g f]
+ C\<^sub>U.dom_char\<^sub>S\<^sub>b\<^sub>C [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
by fastforce
also have "... = C\<^sub>U.P g \<star>\<^sub>C (C\<^sub>U.P f)"
using f g fg by auto
also have "... = C.cod (C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
using f g fg C\<^sub>U.emb.cmp_def [of "(C\<^sub>U.P g, C\<^sub>U.P f)"] by auto
finally show ?thesis by blast
qed
show "C.arr (C\<^sub>U.P (g \<star>\<^sub>C f))"
using f g fg by simp
show "C.dom (C\<^sub>U.P (g \<star>\<^sub>C f)) =
C.cod (C\<^sub>U.\<Phi>\<^sub>P (g, f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
proof -
have "C.dom (C\<^sub>U.P (g \<star>\<^sub>C f)) = C\<^sub>U.P (g \<star>\<^sub>C f)"
using f g fg by simp
also have "... = C.cod (C\<^sub>U.\<Phi>\<^sub>P (g, f))"
- using f g fg C\<^sub>U.\<Phi>\<^sub>P_simps(5) [of g f] C\<^sub>U.cod_char [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
+ using f g fg C\<^sub>U.\<Phi>\<^sub>P_simps(5) [of g f] C\<^sub>U.cod_char\<^sub>S\<^sub>b\<^sub>C [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
by fastforce
also have "... = C.cod (C\<^sub>U.\<Phi>\<^sub>P (g, f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
proof -
have 2: "C\<^sub>U.hcomp (C\<^sub>U.P g) (C\<^sub>U.P f) = C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f"
using f g fg by auto
have "C\<^sub>U.arr (C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f)"
proof -
have "C\<^sub>U.arr (C\<^sub>U.hcomp (C\<^sub>U.P g) (C\<^sub>U.P f))"
using f g fg by simp
thus ?thesis
using 2 by simp
qed
moreover have "C.dom (C\<^sub>U.\<Phi>\<^sub>P (g, f)) = C\<^sub>U.P g \<star>\<^sub>C C\<^sub>U.P f"
- using f g fg 2 C\<^sub>U.VV.arr_char C\<^sub>U.\<Phi>\<^sub>P_simps(4) [of g f]
- C\<^sub>U.dom_char [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
+ using f g fg 2 C\<^sub>U.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.\<Phi>\<^sub>P_simps(4) [of g f]
+ C\<^sub>U.dom_char\<^sub>S\<^sub>b\<^sub>C [of "C\<^sub>U.\<Phi>\<^sub>P (g, f)"]
by fastforce
ultimately have "C.arr (C\<^sub>U.\<Phi>\<^sub>P (g, f) \<cdot>\<^sub>C C\<^sub>U.\<Phi>\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))"
- using f g fg 1 C\<^sub>U.VV.arr_char C\<^sub>U.VV.cod_char C\<^sub>U.hcomp_char
+ using f g fg 1 C\<^sub>U.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.hcomp_char
C\<^sub>U.emb.map_def
by auto
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
qed
thus ?thesis
using D.comp_assoc by auto
qed
also have "... = F (C\<^sub>U.EoP.cmp (g, f)) \<cdot>\<^sub>D \<Phi> (C\<^sub>U.P g, C\<^sub>U.P f)"
- using f g fg C\<^sub>U.EoP.cmp_def C.VV.arr_char C.VV.dom_char C\<^sub>U.emb.map_def
+ using f g fg C\<^sub>U.EoP.cmp_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.emb.map_def
by simp
finally show ?thesis by blast
qed
moreover have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g = F (C\<^sub>U.P g)"
proof -
have "D\<^sub>V.arr (F (C\<^sub>U.P g))"
- using g D\<^sub>V.arr_char C\<^sub>U.P_simps(1) C.ideD(1) F\<^sub>U\<^sub>V.preserves_arr by presburger
+ using g D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.P_simps(1) C.ideD(1) F\<^sub>U\<^sub>V.preserves_arr by presburger
thus ?thesis
using g D\<^sub>V.emb.map_def by simp
qed
moreover have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f = F (C\<^sub>U.P f)"
proof -
have "D\<^sub>V.arr (F (C\<^sub>U.P f))"
- using f D\<^sub>V.arr_char C\<^sub>U.P_simps(1) C.ideD(1) F\<^sub>U\<^sub>V.preserves_arr by presburger
+ using f D\<^sub>V.arr_char\<^sub>S\<^sub>b\<^sub>C C\<^sub>U.P_simps(1) C.ideD(1) F\<^sub>U\<^sub>V.preserves_arr by presburger
thus ?thesis
using f D\<^sub>V.emb.map_def by simp
qed
ultimately show ?thesis
using fg by argo
qed
finally show "(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g,
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0 (src\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, \<tau>\<^sub>0 (src\<^sub>C f)]
= \<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))"
by argo
qed
qed
interpretation EQ: equivalence_of_bicategories_and_pseudonatural_equivalence_left (* 17 sec *)
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1
F \<Phi> \<tau>\<^sub>0 \<tau>\<^sub>1
proof -
interpret E: equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1
using induces_equivalence_of_bicategories by blast
interpret \<tau>: pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp F \<Phi> \<tau>\<^sub>0 \<tau>\<^sub>1
..
show "equivalence_of_bicategories_and_pseudonatural_equivalence_left
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1
EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1
F \<Phi> \<tau>\<^sub>0 \<tau>\<^sub>1"
..
qed
definition right_map
where "right_map \<equiv> EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map"
definition right_cmp
where "right_cmp \<equiv> EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp"
definition unit\<^sub>0
where "unit\<^sub>0 \<equiv> EQ.unit.map\<^sub>0"
definition unit\<^sub>1
where "unit\<^sub>1 \<equiv> EQ.unit.map\<^sub>1"
definition counit\<^sub>0
where "counit\<^sub>0 \<equiv> EQ.counit.map\<^sub>0"
definition counit\<^sub>1
where "counit\<^sub>1 \<equiv> EQ.counit.map\<^sub>1"
theorem extends_to_equivalence_of_bicategories:
shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi> right_map right_cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1"
unfolding right_map_def right_cmp_def unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def
..
end
locale converse_equivalence_pseudofunctor = (* 10 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
begin
interpretation E: equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi>\<^sub>F F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1
using F.extends_to_equivalence_of_bicategories by simp
interpretation E': converse_equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi>\<^sub>F F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1
..
sublocale equivalence_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F.right_map F.right_cmp
using E'.equivalence_pseudofunctor_left by simp
lemma is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F.right_map F.right_cmp"
..
end
(*
* TODO: Change the following definition to be based on the existence of an equivalence of
* bicategories, rather than an equivalence pseudofunctor. This will require a few changes
* elsewhere.
*)
definition equivalent_bicategories
where "equivalent_bicategories V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C \<equiv>
\<exists>F \<Phi>. equivalence_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>"
lemma equivalent_bicategories_reflexive:
assumes "bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C"
shows "equivalent_bicategories V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C"
proof -
interpret bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
using assms by simp
interpret I: identity_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C ..
show ?thesis
using I.equivalence_pseudofunctor_axioms equivalent_bicategories_def by blast
qed
lemma equivalent_bicategories_symmetric:
assumes "equivalent_bicategories V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
shows "equivalent_bicategories V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C"
proof -
obtain F \<Phi>\<^sub>F where F: "equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F"
using assms equivalent_bicategories_def by blast
interpret F: equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F
using F by simp
interpret G: converse_equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F
..
show ?thesis
using G.is_equivalence_pseudofunctor equivalent_bicategories_def by blast
qed
lemma equivalent_bicategories_transitive:
assumes "equivalent_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C"
and "equivalent_bicategories V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
shows "equivalent_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
proof -
obtain F \<Phi>\<^sub>F where F: "equivalence_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F"
using assms(1) equivalent_bicategories_def by blast
obtain G \<Phi>\<^sub>G where G: "equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G"
using assms(2) equivalent_bicategories_def by blast
interpret F: equivalence_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F
using F by simp
interpret G: equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
using G by simp
interpret GF: composite_equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G ..
show "equivalent_bicategories V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
using equivalent_bicategories_def GF.equivalence_pseudofunctor_axioms by blast
qed
end
diff --git a/thys/Bicategory/InternalAdjunction.thy b/thys/Bicategory/InternalAdjunction.thy
--- a/thys/Bicategory/InternalAdjunction.thy
+++ b/thys/Bicategory/InternalAdjunction.thy
@@ -1,3342 +1,3342 @@
(* Title: InternalAdjunction
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Adjunctions in a Bicategory"
theory InternalAdjunction
imports CanonicalIsos Strictness
begin
text \<open>
An \emph{internal adjunction} in a bicategory is a four-tuple \<open>(f, g, \<eta>, \<epsilon>)\<close>,
where \<open>f\<close> and \<open>g\<close> are antiparallel 1-cells and \<open>\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>\<close> and
\<open>\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> src g\<guillemotright>\<close> are 2-cells, such that the familiar ``triangle''
(or ``zig-zag'') identities are satisfied. We state the triangle identities
in two equivalent forms, each of which is convenient in certain situations.
\<close>
locale adjunction_in_bicategory =
adjunction_data_in_bicategory +
assumes triangle_left: "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
and triangle_right: "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
begin
lemma triangle_left':
shows "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] = f"
using triangle_left triangle_equiv_form by simp
lemma triangle_right':
shows "\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] = g"
using triangle_right triangle_equiv_form by simp
end
text \<open>
Internal adjunctions have a number of properties, which we now develop,
that generalize those of ordinary adjunctions involving functors and
natural transformations.
\<close>
context bicategory
begin
lemma adjunction_unit_determines_counit:
assumes "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>"
and "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>'"
shows "\<epsilon> = \<epsilon>'"
proof -
interpret E: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms(1) by auto
interpret E': adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using assms(2) by auto
text \<open>
Note that since we want to prove the the result for an arbitrary bicategory,
not just in for a strict bicategory, the calculation is a little more involved
than one might expect from a treatment that suppresses canonical isomorphisms.
\<close>
have "\<epsilon> = \<epsilon> \<cdot> (f \<star> \<r>[g] \<cdot> (g \<star> \<epsilon>') \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g])"
using E'.triangle_right' comp_arr_dom by simp
also have "... = \<epsilon> \<cdot> (f \<star> \<r>[g]) \<cdot> (f \<star> g \<star> \<epsilon>') \<cdot> (f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar whisker_left by simp
also have "... = \<epsilon> \<cdot> ((f \<star> \<r>[g]) \<cdot> (f \<star> g \<star> \<epsilon>')) \<cdot> (f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using comp_assoc by simp
also have "... = \<epsilon> \<cdot> \<r>[f \<star> g] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, src g] \<cdot> (f \<star> g \<star> \<epsilon>')) \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "f \<star> \<r>[g] = \<r>[f \<star> g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, src g]"
using E.antipar(1) runit_hcomp(3) by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = (\<epsilon> \<cdot> \<r>[f \<star> g]) \<cdot> ((f \<star> g) \<star> \<epsilon>') \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar E'.counit_in_hom assoc'_naturality [of f g \<epsilon>'] comp_assoc by simp
also have "... = \<r>[trg f] \<cdot> ((\<epsilon> \<star> trg f) \<cdot> ((f \<star> g) \<star> \<epsilon>')) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar E.counit_in_hom runit_naturality [of \<epsilon>] comp_assoc by simp
also have "... = (\<l>[src g] \<cdot> (src g \<star> \<epsilon>')) \<cdot> (\<epsilon> \<star> f \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "(\<epsilon> \<star> trg f) \<cdot> ((f \<star> g) \<star> \<epsilon>') = (src g \<star> \<epsilon>') \<cdot> (\<epsilon> \<star> f \<star> g)"
using E.antipar interchange E.counit_in_hom comp_arr_dom comp_cod_arr
by (metis E'.counit_simps(1-3) E.counit_simps(1-3))
thus ?thesis
using E.antipar comp_assoc unitor_coincidence by simp
qed
also have "... = \<epsilon>' \<cdot> \<l>[f \<star> g] \<cdot> (\<epsilon> \<star> f \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "\<l>[src g] \<cdot> (src g \<star> \<epsilon>') = \<epsilon>' \<cdot> \<l>[f \<star> g]"
using E.antipar lunit_naturality [of \<epsilon>'] by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[trg f, f, g] \<cdot> (\<epsilon> \<star> f \<star> g)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar lunit_hcomp comp_assoc by simp
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f \<star> g, f, g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g])) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar assoc'_naturality [of \<epsilon> f g] comp_assoc by simp
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> f, g] \<cdot> (f \<star> \<eta> \<star> g)) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "\<a>\<^sup>-\<^sup>1[f \<star> g, f, g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot> (f \<star> \<a>[g, f, g]) =
(\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, g]"
using E.antipar iso_assoc' pentagon' comp_assoc
invert_side_of_triangle(2)
[of "\<a>\<^sup>-\<^sup>1[f \<star> g, f, g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g]"
"(\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, g]" "f \<star> \<a>\<^sup>-\<^sup>1[g, f, g]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot>
((f \<star> \<eta>) \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg g, g] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar assoc'_naturality [of f \<eta> g] comp_assoc by simp
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot>
((f \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[f] \<star> g)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, trg g, g] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g]) = \<r>\<^sup>-\<^sup>1[f] \<star> g"
proof -
have "\<r>\<^sup>-\<^sup>1[f] \<star> g = inv (\<r>[f] \<star> g)"
using E.antipar by simp
also have "... = inv ((f \<star> \<l>[g]) \<cdot> \<a>[f, trg g, g])"
using E.antipar by (simp add: triangle)
also have "... = \<a>\<^sup>-\<^sup>1[f, trg g, g] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar inv_comp by simp
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] \<star> g)"
using E.antipar whisker_right by simp
also have "... = \<epsilon>'"
using E.triangle_left' comp_arr_dom by simp
finally show ?thesis by simp
qed
end
subsection "Adjoint Transpose"
context adjunction_in_bicategory
begin
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
text \<open>
Just as for an ordinary adjunction between categories, an adjunction in a bicategory
determines bijections between hom-sets. There are two versions of this relationship:
depending on whether the transposition is occurring on the left (\emph{i.e.}~``output'')
side or the right (\emph{i.e.}~``input'') side.
\<close>
definition trnl\<^sub>\<eta>
where "trnl\<^sub>\<eta> v \<mu> \<equiv> (g \<star> \<mu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
definition trnl\<^sub>\<epsilon>
where "trnl\<^sub>\<epsilon> u \<nu> \<equiv> \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> \<nu>)"
lemma adjoint_transpose_left:
assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u"
shows "trnl\<^sub>\<eta> v \<in> hom (f \<star> v) u \<rightarrow> hom v (g \<star> u)"
and "trnl\<^sub>\<epsilon> u \<in> hom v (g \<star> u) \<rightarrow> hom (f \<star> v) u"
and "\<guillemotleft>\<mu> : f \<star> v \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) = \<mu>"
and "\<guillemotleft>\<nu> : v \<Rightarrow> g \<star> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) = \<nu>"
and "bij_betw (trnl\<^sub>\<eta> v) (hom (f \<star> v) u) (hom v (g \<star> u))"
and "bij_betw (trnl\<^sub>\<epsilon> u) (hom v (g \<star> u)) (hom (f \<star> v) u)"
proof -
show A: "trnl\<^sub>\<eta> v \<in> hom (f \<star> v) u \<rightarrow> hom v (g \<star> u)"
using assms antipar trnl\<^sub>\<eta>_def by fastforce
show B: "trnl\<^sub>\<epsilon> u \<in> hom v (g \<star> u) \<rightarrow> hom (f \<star> v) u"
using assms antipar trnl\<^sub>\<epsilon>_def by fastforce
show C: "\<And>\<mu>. \<guillemotleft>\<mu> : f \<star> v \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) = \<mu>"
proof -
fix \<mu>
assume \<mu>: "\<guillemotleft>\<mu> : f \<star> v \<Rightarrow> u\<guillemotright>"
have "trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) =
\<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> (g \<star> \<mu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v])"
using trnl\<^sub>\<eta>_def trnl\<^sub>\<epsilon>_def by simp
also have "... = \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> g \<star> \<mu>)) \<cdot> (f \<star> \<a>[g, f, v]) \<cdot>
(f \<star> \<eta> \<star> v) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
using assms \<mu> antipar whisker_left comp_assoc by auto
also have "... = \<l>[u] \<cdot> ((\<epsilon> \<star> u) \<cdot> ((f \<star> g) \<star> \<mu>)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot>
(f \<star> \<eta> \<star> v) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
using assms \<mu> antipar assoc'_naturality [of f g \<mu>] comp_assoc by fastforce
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot>
(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot>
(f \<star> \<eta> \<star> v) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
proof -
have "(\<epsilon> \<star> u) \<cdot> ((f \<star> g) \<star> \<mu>) = (trg u \<star> \<mu>) \<cdot> (\<epsilon> \<star> f \<star> v)"
using assms \<mu> antipar comp_cod_arr comp_arr_dom
interchange [of "trg u" \<epsilon> \<mu> "f \<star> v"] interchange [of \<epsilon> "f \<star> g" u \<mu>]
by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot> \<a>[trg f, f, v] \<cdot>
((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<star> v) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
proof -
have 1: "(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot> (f \<star> \<eta> \<star> v) =
\<a>[trg f, f, v] \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg v, v]"
proof -
have "(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot> (f \<star> \<eta> \<star> v) =
(\<epsilon> \<star> f \<star> v) \<cdot>
\<a>[f \<star> g, f, v] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, v] \<cdot>
(f \<star> \<eta> \<star> v)"
proof -
have "(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot> (f \<star> \<eta> \<star> v) =
(\<epsilon> \<star> f \<star> v) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v])) \<cdot> (f \<star> \<eta> \<star> v)"
using comp_assoc by simp
also have "... = (\<epsilon> \<star> f \<star> v) \<cdot>
\<a>[f \<star> g, f, v] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, v] \<cdot>
(f \<star> \<eta> \<star> v)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) =
\<a>[f \<star> g, f, v] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, v]"
using assms antipar canI_associator_0 whisker_can_left_0 whisker_can_right_0
canI_associator_hcomp(1-3)
by simp
thus ?thesis
using comp_assoc by simp
qed
finally show ?thesis by blast
qed
also have "... = ((\<epsilon> \<star> f \<star> v) \<cdot> \<a>[f \<star> g, f, v]) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> ((f \<star> \<eta>) \<star> v) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v]"
using assms \<mu> antipar assoc'_naturality [of f \<eta> v] comp_assoc by simp
also have "... = (\<a>[trg f, f, v] \<cdot> ((\<epsilon> \<star> f) \<star> v)) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> ((f \<star> \<eta>) \<star> v) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v]"
using assms \<mu> antipar assoc_naturality [of \<epsilon> f v] by simp
also have "... = \<a>[trg f, f, v] \<cdot>
(((\<epsilon> \<star> f) \<star> v) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> ((f \<star> \<eta>) \<star> v)) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v]"
using comp_assoc by simp
also have "... = \<a>[trg f, f, v] \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg v, v]"
using assms \<mu> antipar whisker_right by simp
finally show ?thesis by simp
qed
show ?thesis
using 1 comp_assoc by metis
qed
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot>
\<a>[trg f, f, v] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg v, v] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
using assms \<mu> antipar triangle_left by simp
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot> can (\<^bold>\<langle>trg u\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>v\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>v\<^bold>\<rangle>)"
using assms \<mu> antipar canI_unitor_0 canI_associator_1
canI_associator_1(1-2) [of f v] whisker_can_right_0 whisker_can_left_0
by simp
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> v]"
unfolding can_def using assms antipar comp_arr_dom comp_cod_arr \<ll>_ide_simp
by simp
also have "... = (\<l>[u] \<cdot> \<l>\<^sup>-\<^sup>1[u]) \<cdot> \<mu>"
using assms \<mu> lunit'_naturality [of \<mu>] comp_assoc by auto
also have "... = \<mu>"
using assms \<mu> comp_cod_arr iso_lunit comp_arr_inv inv_is_inverse by auto
finally show "trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) = \<mu>" by simp
qed
show D: "\<And>\<nu>. \<guillemotleft>\<nu> : v \<Rightarrow> g \<star> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) = \<nu>"
proof -
fix \<nu>
assume \<nu>: "\<guillemotleft>\<nu> : v \<Rightarrow> g \<star> u\<guillemotright>"
have "trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) =
(g \<star> \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> \<nu>)) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
using trnl\<^sub>\<eta>_def trnl\<^sub>\<epsilon>_def by simp
also have "... = (g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> (g \<star> f \<star> \<nu>) \<cdot>
\<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
using assms \<nu> antipar interchange [of g "g \<cdot> g \<cdot> g"] comp_assoc by auto
also have "... = ((g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot>
\<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u)) \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> f \<star> \<nu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v] =
\<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> f \<star> \<nu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v] =
\<a>[g, f, g \<star> u] \<cdot> ((g \<star> f) \<star> \<nu>) \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> f \<star> \<nu>) \<cdot> \<a>[g, f, v] = \<a>[g, f, g \<star> u] \<cdot> ((g \<star> f) \<star> \<nu>)"
using assms \<nu> antipar assoc_naturality [of g f \<nu>] by auto
thus ?thesis
using assms comp_assoc by metis
qed
also have "... = \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "((g \<star> f) \<star> \<nu>) \<cdot> (\<eta> \<star> v) = (\<eta> \<star> g \<star> u) \<cdot> (trg v \<star> \<nu>)"
using assms \<nu> antipar comp_arr_dom comp_cod_arr
interchange [of "g \<star> f" \<eta> \<nu> v] interchange [of \<eta> "trg v" "g \<star> u" \<nu>]
by auto
thus ?thesis
using comp_assoc by metis
qed
finally show ?thesis by blast
qed
thus ?thesis using comp_assoc by simp
qed
also have "... = \<l>[g \<star> u] \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) =
\<l>[g \<star> u]"
proof -
have "(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) =
(g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot>
((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<star> u) \<cdot>
\<a>\<^sup>-\<^sup>1[trg v, g, u]"
proof -
have "(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) =
(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot>
((\<eta> \<star> g \<star> u) \<cdot> \<a>[trg v, g, u]) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using assms antipar comp_arr_dom comp_assoc comp_assoc_assoc'(1) by simp
also have "... = (g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot>
(\<a>[g \<star> f, g, u] \<cdot> ((\<eta> \<star> g) \<star> u)) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using assms antipar assoc_naturality [of \<eta> g u] by simp
also have "... = (g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot>
((g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using comp_assoc by simp
also have "... = (g \<star> \<l>[u]) \<cdot> ((\<a>[g, trg u, u] \<cdot> \<a>\<^sup>-\<^sup>1[g, trg u, u]) \<cdot> (g \<star> \<epsilon> \<star> u)) \<cdot>
((g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
proof -
have "(\<a>[g, trg u, u] \<cdot> \<a>\<^sup>-\<^sup>1[g, trg u, u]) \<cdot> (g \<star> \<epsilon> \<star> u) = g \<star> \<epsilon> \<star> u"
using assms antipar comp_cod_arr comp_assoc_assoc'(1) by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<a>\<^sup>-\<^sup>1[g, trg u, u] \<cdot> (g \<star> \<epsilon> \<star> u)) \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u] \<cdot>
((\<eta> \<star> g) \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using comp_assoc by simp
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (((g \<star> \<epsilon>) \<star> u) \<cdot> (\<a>\<^sup>-\<^sup>1[g, f \<star> g, u] \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u)) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using assms antipar assoc'_naturality [of g \<epsilon> u] comp_assoc by simp
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot>
((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<star> u) \<cdot>
\<a>\<^sup>-\<^sup>1[trg v, g, u]"
proof -
have "\<a>\<^sup>-\<^sup>1[g, f \<star> g, u] \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u] =
\<a>[g, f, g] \<star> u"
using assms antipar canI_associator_0 whisker_can_left_0 whisker_can_right_0
canI_associator_hcomp
by simp
hence "((g \<star> \<epsilon>) \<star> u) \<cdot>
(\<a>\<^sup>-\<^sup>1[g, f \<star> g, u] \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u) =
(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<star> u"
using assms antipar whisker_right by simp
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg g, g, u]"
using assms antipar triangle_right by simp
also have "... = can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>) (\<^bold>\<langle>trg g\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>)"
proof -
have "(g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg g, g, u] =
((g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg g, g, u])"
using comp_assoc by simp
moreover have "... = can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>) (\<^bold>\<langle>trg g\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>)"
using assms antipar canI_unitor_0 canI_associator_1 [of g u] inv_can
whisker_can_left_0 whisker_can_right_0
by simp
ultimately show ?thesis by simp
qed
also have "... = \<l>[g \<star> u]"
unfolding can_def using assms comp_arr_dom comp_cod_arr \<ll>_ide_simp by simp
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = (\<l>[g \<star> u] \<cdot> \<l>\<^sup>-\<^sup>1[g \<star> u]) \<cdot> \<nu>"
using assms \<nu> lunit'_naturality comp_assoc by auto
also have "... = \<nu>"
using assms \<nu> comp_cod_arr iso_lunit comp_arr_inv inv_is_inverse by auto
finally show "trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) = \<nu>" by simp
qed
show "bij_betw (trnl\<^sub>\<eta> v) (hom (f \<star> v) u) (hom v (g \<star> u))"
using A B C D by (intro bij_betwI) auto
show "bij_betw (trnl\<^sub>\<epsilon> u) (hom v (g \<star> u)) (hom (f \<star> v) u)"
using A B C D by (intro bij_betwI) auto
qed
lemma trnl\<^sub>\<epsilon>_comp:
assumes "ide u" and "seq \<mu> \<nu>" and "src f = trg \<mu>"
shows "trnl\<^sub>\<epsilon> u (\<mu> \<cdot> \<nu>) = trnl\<^sub>\<epsilon> u \<mu> \<cdot> (f \<star> \<nu>)"
using assms trnl\<^sub>\<epsilon>_def whisker_left [of f \<mu> \<nu>] comp_assoc by auto
definition trnr\<^sub>\<eta>
where "trnr\<^sub>\<eta> v \<mu> \<equiv> (\<mu> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
definition trnr\<^sub>\<epsilon>
where "trnr\<^sub>\<epsilon> u \<nu> \<equiv> \<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> (\<nu> \<star> g)"
lemma adjoint_transpose_right:
assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f"
shows "trnr\<^sub>\<eta> v \<in> hom (v \<star> g) u \<rightarrow> hom v (u \<star> f)"
and "trnr\<^sub>\<epsilon> u \<in> hom v (u \<star> f) \<rightarrow> hom (v \<star> g) u"
and "\<guillemotleft>\<mu> : v \<star> g \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) = \<mu>"
and "\<guillemotleft>\<nu> : v \<Rightarrow> u \<star> f\<guillemotright> \<Longrightarrow> trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) = \<nu>"
and "bij_betw (trnr\<^sub>\<eta> v) (hom (v \<star> g) u) (hom v (u \<star> f))"
and "bij_betw (trnr\<^sub>\<epsilon> u) (hom v (u \<star> f)) (hom (v \<star> g) u)"
proof -
show A: "trnr\<^sub>\<eta> v \<in> hom (v \<star> g) u \<rightarrow> hom v (u \<star> f)"
using assms antipar trnr\<^sub>\<eta>_def by fastforce
show B: "trnr\<^sub>\<epsilon> u \<in> hom v (u \<star> f) \<rightarrow> hom (v \<star> g) u"
using assms antipar trnr\<^sub>\<epsilon>_def by fastforce
show C: "\<And>\<mu>. \<guillemotleft>\<mu> : v \<star> g \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) = \<mu>"
proof -
fix \<mu>
assume \<mu>: "\<guillemotleft>\<mu> : v \<star> g \<Rightarrow> u\<guillemotright>"
have "trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) =
\<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> ((\<mu> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v] \<star> g)"
unfolding trnr\<^sub>\<epsilon>_def trnr\<^sub>\<eta>_def by simp
also have "... = \<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> (\<a>[u, f, g] \<cdot> ((\<mu> \<star> f) \<star> g)) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using assms \<mu> antipar whisker_right comp_assoc by auto
also have "... = \<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> ((\<mu> \<star> f \<star> g) \<cdot> \<a>[v \<star> g, f, g]) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using assms \<mu> antipar assoc_naturality [of \<mu> f g] by auto
also have "... = \<r>[u] \<cdot> ((u \<star> \<epsilon>) \<cdot> (\<mu> \<star> f \<star> g)) \<cdot> \<a>[v \<star> g, f, g] \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using comp_assoc by auto
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot> ((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
proof -
have "(u \<star> \<epsilon>) \<cdot> (\<mu> \<star> f \<star> g) = (\<mu> \<star> src u) \<cdot> ((v \<star> g) \<star> \<epsilon>)"
using assms \<mu> antipar comp_arr_dom comp_cod_arr
interchange [of \<mu> "v \<star> g" "src u" \<epsilon>] interchange [of u \<mu> \<epsilon> "f \<star> g"]
by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot>
(((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g)) \<cdot>
(\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using comp_assoc by simp
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot>
\<a>[v, src v, g]) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
proof -
have "((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) =
\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot> \<a>[v, src v, g]"
proof -
have "((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) =
((\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> \<a>[v, g, src u]) \<cdot> ((v \<star> g) \<star> \<epsilon>)) \<cdot>
\<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g)"
proof -
have "arr v \<and> dom v = v \<and> cod v = v"
using assms(2) ide_char by blast
moreover have "arr g \<and> dom g = g \<and> cod g = g"
using ide_right ide_char by blast
ultimately show ?thesis
by (metis (no_types) antipar(2) assms(3-4) assoc_naturality
counit_simps(1,3,5) hcomp_reassoc(1) comp_assoc)
qed
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (\<a>[v, g, src u] \<cdot> ((v \<star> g) \<star> \<epsilon>)) \<cdot>
\<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g)"
using comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> ((v \<star> g \<star> \<epsilon>) \<cdot> \<a>[v, g, f \<star> g]) \<cdot>
\<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g \<star> f, g] \<cdot> \<a>[v, g \<star> f, g]) \<cdot> ((v \<star> \<eta>) \<star> g)"
proof -
have "\<a>[v, g, src u] \<cdot> ((v \<star> g) \<star> \<epsilon>) = (v \<star> g \<star> \<epsilon>) \<cdot> \<a>[v, g, f \<star> g]"
using assms antipar assoc_naturality [of v g \<epsilon>] by simp
moreover have "(\<a>\<^sup>-\<^sup>1[v, g \<star> f, g] \<cdot> \<a>[v, g \<star> f, g]) \<cdot> ((v \<star> \<eta>) \<star> g) = (v \<star> \<eta>) \<star> g"
using assms antipar comp_cod_arr comp_assoc_assoc'(2) by simp
ultimately show ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> g \<star> \<epsilon>) \<cdot>
\<a>[v, g, f \<star> g] \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot>
\<a>\<^sup>-\<^sup>1[v, g \<star> f, g] \<cdot> \<a>[v, g \<star> f, g] \<cdot> ((v \<star> \<eta>) \<star> g)"
using comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> ((v \<star> g \<star> \<epsilon>) \<cdot>
(\<a>[v, g, f \<star> g] \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot>
\<a>\<^sup>-\<^sup>1[v, g \<star> f, g]) \<cdot> (v \<star> \<eta> \<star> g)) \<cdot> \<a>[v, src v, g]"
using assms antipar assoc_naturality [of v \<eta> g] comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot>
((v \<star> g \<star> \<epsilon>) \<cdot> (v \<star> \<a>[g, f, g]) \<cdot> (v \<star> \<eta> \<star> g)) \<cdot>
\<a>[v, src v, g]"
proof -
have "\<a>[v, g, f \<star> g] \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[v, g \<star> f, g] =
v \<star> \<a>[g, f, g]"
using assms antipar canI_associator_0 canI_associator_hcomp
whisker_can_left_0 whisker_can_right_0
by simp
thus ?thesis
using assms antipar whisker_left by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot>
(v \<star> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot>
\<a>[v, src v, g]"
using assms antipar whisker_left by simp
finally show ?thesis by simp
qed
thus ?thesis by auto
qed
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot>
\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot>
\<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using triangle_right comp_assoc by simp
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot> \<r>\<^sup>-\<^sup>1[v \<star> g]"
proof -
have "\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g) = \<r>\<^sup>-\<^sup>1[v \<star> g]"
proof -
have "\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g) =
\<a>\<^sup>-\<^sup>1[v, g, trg f] \<cdot> can (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>src g\<^bold>\<rangle>\<^sub>0) (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>)"
using assms canI_unitor_0 canI_associator_1(2-3) whisker_can_left_0(1)
whisker_can_right_0
by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src g] \<cdot> can (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>src g\<^bold>\<rangle>\<^sub>0) (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>)"
using antipar by simp
(* TODO: There should be an alternate version of whisker_can_left for this. *)
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src g] \<cdot> (v \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>src g\<^bold>\<rangle>\<^sub>0) \<^bold>\<langle>g\<^bold>\<rangle>)"
using assms canI_unitor_0(2) whisker_can_left_0 by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src g] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g])"
using assms canI_unitor_0(2) by simp
also have "... = \<r>\<^sup>-\<^sup>1[v \<star> g]"
using assms runit_hcomp(2) by simp
finally have "\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g) =
\<r>\<^sup>-\<^sup>1[v \<star> g]"
by simp
thus ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = (\<r>[u] \<cdot> \<r>\<^sup>-\<^sup>1[u]) \<cdot> \<mu>"
using assms \<mu> runit'_naturality [of \<mu>] comp_assoc by auto
also have "... = \<mu>"
using \<mu> comp_cod_arr iso_runit inv_is_inverse comp_arr_inv by auto
finally show "trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) = \<mu>" by simp
qed
show D: "\<And>\<nu>. \<guillemotleft>\<nu> : v \<Rightarrow> u \<star> f\<guillemotright> \<Longrightarrow> trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) = \<nu>"
proof -
fix \<nu>
assume \<nu>: "\<guillemotleft>\<nu> : v \<Rightarrow> u \<star> f\<guillemotright>"
have "trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) =
(\<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> (\<nu> \<star> g) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
unfolding trnr\<^sub>\<eta>_def trnr\<^sub>\<epsilon>_def by simp
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
(((\<nu> \<star> g) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f]) \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms \<nu> antipar whisker_right comp_assoc by auto
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
(\<a>\<^sup>-\<^sup>1[u \<star> f, g, f] \<cdot> (\<nu> \<star> g \<star> f)) \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms \<nu> antipar assoc'_naturality [of \<nu> g f] by auto
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> f, g, f] \<cdot> ((\<nu> \<star> g \<star> f) \<cdot> (v \<star> \<eta>)) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> f, g, f] \<cdot> (((u \<star> f) \<star> \<eta>) \<cdot> (\<nu> \<star> src v)) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
proof -
have "(\<nu> \<star> g \<star> f) \<cdot> (v \<star> \<eta>) = ((u \<star> f) \<star> \<eta>) \<cdot> (\<nu> \<star> src v)"
using assms \<nu> antipar interchange [of "u \<star> f" \<nu> \<eta> "src v"]
interchange [of \<nu> v "g \<star> f" \<eta>] comp_arr_dom comp_cod_arr
by auto
thus ?thesis by simp
qed
also have "... = ((\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot>
((\<a>[u, f, g] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u \<star> f, g, f]) \<cdot>
((u \<star> f) \<star> \<eta>)) \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = ((\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot>
(\<a>\<^sup>-\<^sup>1[u, f \<star> g, f] \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> \<a>[u, f, g \<star> f]) \<cdot>
((u \<star> f) \<star> \<eta>)) \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar canI_associator_hcomp canI_associator_0 whisker_can_left_0
whisker_can_right_0
by simp
also have "... = ((\<r>[u] \<star> f) \<cdot> (((u \<star> \<epsilon>) \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[u, f \<star> g, f]) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> (\<a>[u, f, g \<star> f]) \<cdot>
((u \<star> f) \<star> \<eta>)) \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = ((\<r>[u] \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot> (u \<star> \<epsilon> \<star> f)) \<cdot>
(u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> ((u \<star> f \<star> \<eta>) \<cdot> \<a>[u, f, src f])) \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar assoc'_naturality [of u \<epsilon> f] assoc_naturality [of u f \<eta>]
by auto
also have "... = (\<r>[u] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot>
((u \<star> \<epsilon> \<star> f) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> (u \<star> f \<star> \<eta>)) \<cdot> \<a>[u, f, src f] \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = (\<r>[u] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot>
(u \<star> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) \<cdot> \<a>[u, f, src f] \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar whisker_left by auto
also have "... = ((\<r>[u] \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot> (u \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<a>[u, f, src f])) \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar triangle_left comp_assoc by simp
also have "... = \<r>[u \<star> f] \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
proof -
have "(\<r>[u] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot> (u \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<a>[u, f, src f] =
((u \<star> \<l>[f]) \<cdot> (\<a>[u, src u, f] \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f])) \<cdot>
(u \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<a>[u, f, src f]"
using assms ide_left ide_right antipar triangle comp_assoc by metis
also have "... = (u \<star> \<r>[f]) \<cdot> \<a>[u, f, src f]"
using assms antipar canI_associator_1 canI_unitor_0 whisker_can_left_0
whisker_can_right_0 canI_associator_1
by simp
also have "... = \<r>[u \<star> f]"
using assms antipar runit_hcomp by simp
finally show ?thesis by simp
qed
also have "... = (\<r>[u \<star> f] \<cdot> \<r>\<^sup>-\<^sup>1[u \<star> f]) \<cdot> \<nu>"
using assms \<nu> runit'_naturality [of \<nu>] comp_assoc by auto
also have "... = \<nu>"
using assms \<nu> comp_cod_arr comp_arr_inv inv_is_inverse iso_runit by auto
finally show "trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) = \<nu>" by auto
qed
show "bij_betw (trnr\<^sub>\<eta> v) (hom (v \<star> g) u) (hom v (u \<star> f))"
using A B C D by (intro bij_betwI, auto)
show "bij_betw (trnr\<^sub>\<epsilon> u) (hom v (u \<star> f)) (hom (v \<star> g) u)"
using A B C D by (intro bij_betwI, auto)
qed
lemma trnr\<^sub>\<eta>_comp:
assumes "ide v" and "seq \<mu> \<nu>" and "src \<mu> = trg f"
shows "trnr\<^sub>\<eta> v (\<mu> \<cdot> \<nu>) = (\<mu> \<star> f) \<cdot> trnr\<^sub>\<eta> v \<nu>"
using assms trnr\<^sub>\<eta>_def whisker_right comp_assoc by auto
end
text \<open>
It is useful to have at hand the simpler versions of the preceding results that
hold in a normal bicategory and in a strict bicategory.
\<close>
locale adjunction_in_normal_bicategory =
normal_bicategory +
adjunction_in_bicategory
begin
lemma triangle_left:
shows "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = f"
using triangle_left strict_lunit strict_runit by simp
lemma triangle_right:
shows "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = g"
using triangle_right strict_lunit strict_runit by simp
lemma trnr\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<mu> \<in> hom (v \<star> g) u"
shows "trnr\<^sub>\<eta> v \<mu> = (\<mu> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>)"
unfolding trnr\<^sub>\<eta>_def
using assms antipar strict_runit' comp_arr_ide [of "\<r>\<^sup>-\<^sup>1[v]" "v \<star> \<eta>"] hcomp_arr_obj
by auto
lemma trnr\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<nu> \<in> hom v (u \<star> f)"
shows "trnr\<^sub>\<epsilon> u \<nu> = (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> (\<nu> \<star> g)"
unfolding trnr\<^sub>\<epsilon>_def
using assms antipar strict_runit comp_ide_arr hcomp_arr_obj by auto
lemma trnl\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<mu> \<in> hom (f \<star> v) u"
shows "trnl\<^sub>\<eta> v \<mu> = (g \<star> \<mu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v)"
using assms trnl\<^sub>\<eta>_def antipar strict_lunit comp_arr_dom hcomp_obj_arr by auto
lemma trnl\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<nu> \<in> hom v (g \<star> u)"
shows "trnl\<^sub>\<epsilon> u \<nu> = (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> \<nu>)"
using assms trnl\<^sub>\<epsilon>_def antipar strict_lunit comp_cod_arr hcomp_obj_arr by auto
end
locale adjunction_in_strict_bicategory =
strict_bicategory +
adjunction_in_normal_bicategory
begin
lemma triangle_left:
shows "(\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f"
using ide_left ide_right antipar triangle_left strict_assoc' comp_cod_arr
by (metis dom_eqI ideD(1) seqE)
lemma triangle_right:
shows "(g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) = g"
using ide_left ide_right antipar triangle_right strict_assoc comp_cod_arr
by (metis ideD(1) ideD(2) seqE)
lemma trnr\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<mu> \<in> hom (v \<star> g) u"
shows "trnr\<^sub>\<eta> v \<mu> = (\<mu> \<star> f) \<cdot> (v \<star> \<eta>)"
using assms antipar trnr\<^sub>\<eta>_eq strict_assoc' comp_ide_arr [of "\<a>\<^sup>-\<^sup>1[v, g, f]" "v \<star> \<eta>"]
by force
lemma trnr\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<nu> \<in> hom v (u \<star> f)"
shows "trnr\<^sub>\<epsilon> u \<nu> = (u \<star> \<epsilon>) \<cdot> (\<nu> \<star> g)"
using assms antipar trnr\<^sub>\<epsilon>_eq strict_assoc comp_ide_arr [of "\<a>[u, f, g]" "\<nu> \<star> g"]
by force
lemma trnl\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<mu> \<in> hom (f \<star> v) u"
shows "trnl\<^sub>\<eta> v \<mu> = (g \<star> \<mu>) \<cdot> (\<eta> \<star> v)"
using assms antipar trnl\<^sub>\<eta>_eq strict_assoc comp_ide_arr [of "\<a>[g, f, v]" "\<eta> \<star> v"]
by force
lemma trnl\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<nu> \<in> hom v (g \<star> u)"
shows "trnl\<^sub>\<epsilon> u \<nu> = (\<epsilon> \<star> u) \<cdot> (f \<star> \<nu>)"
using assms antipar trnl\<^sub>\<epsilon>_eq strict_assoc' comp_ide_arr [of "\<a>\<^sup>-\<^sup>1[f, g, u]" "f \<star> \<nu>"]
by fastforce
end
subsection "Preservation Properties for Adjunctions"
text \<open>
Here we show that adjunctions are preserved under isomorphisms of the
left and right adjoints.
\<close>
context bicategory
begin
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
definition adjoint_pair
where "adjoint_pair f g \<equiv> \<exists>\<eta> \<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
(* These would normally be called "maps", but that name is too heavily used already. *)
abbreviation is_left_adjoint
where "is_left_adjoint f \<equiv> \<exists>g. adjoint_pair f g"
abbreviation is_right_adjoint
where "is_right_adjoint g \<equiv> \<exists>f. adjoint_pair f g"
lemma adjoint_pair_antipar:
assumes "adjoint_pair f g"
shows "ide f" and "ide g" and "src f = trg g" and "src g = trg f"
proof -
obtain \<eta> \<epsilon> where A: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using A by auto
show "ide f" by simp
show "ide g" by simp
show "src f = trg g" using A.antipar by simp
show "src g = trg f" using A.antipar by simp
qed
lemma left_adjoint_is_ide:
assumes "is_left_adjoint f"
shows "ide f"
using assms adjoint_pair_antipar by auto
lemma right_adjoint_is_ide:
assumes "is_right_adjoint f"
shows "ide f"
using assms adjoint_pair_antipar by auto
lemma adjunction_preserved_by_iso_right:
assumes "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "adjunction_in_bicategory V H \<a> \<i> src trg f g' ((\<phi> \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (f \<star> inv \<phi>))"
proof
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by auto
show "ide f" by simp
show "ide g'"
using assms(2) isomorphic_def by auto
show "\<guillemotleft>(\<phi> \<star> f) \<cdot> \<eta> : src f \<Rightarrow> g' \<star> f\<guillemotright>"
using assms A.antipar by fastforce
show "\<guillemotleft>\<epsilon> \<cdot> (f \<star> inv \<phi>) : f \<star> g' \<Rightarrow> src g'\<guillemotright>"
using assms A.antipar A.counit_in_hom by auto
show "(\<epsilon> \<cdot> (f \<star> inv \<phi>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g', f] \<cdot> (f \<star> (\<phi> \<star> f) \<cdot> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "(\<epsilon> \<cdot> (f \<star> inv \<phi>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g', f] \<cdot> (f \<star> (\<phi> \<star> f) \<cdot> \<eta>) =
(\<epsilon> \<star> f) \<cdot> (((f \<star> inv \<phi>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g', f]) \<cdot> (f \<star> \<phi> \<star> f) \<cdot> (f \<star> \<eta>)"
using assms A.antipar whisker_right whisker_left comp_assoc by auto
also have "... = (\<epsilon> \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> inv \<phi> \<star> f)) \<cdot> (f \<star> \<phi> \<star> f) \<cdot> (f \<star> \<eta>)"
using assms A.antipar assoc'_naturality [of f "inv \<phi>" f] by auto
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> ((f \<star> inv \<phi> \<star> f) \<cdot> (f \<star> \<phi> \<star> f)) \<cdot> (f \<star> \<eta>)"
using comp_assoc by simp
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> g \<star> f) \<cdot> (f \<star> \<eta>)"
using assms A.antipar comp_inv_arr inv_is_inverse whisker_left
whisker_right [of f "inv \<phi>" \<phi>]
by auto
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)"
using assms A.antipar comp_cod_arr by simp
also have "... = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
using A.triangle_left by simp
finally show ?thesis by simp
qed
show "(g' \<star> \<epsilon> \<cdot> (f \<star> inv \<phi>)) \<cdot> \<a>[g', f, g'] \<cdot> ((\<phi> \<star> f) \<cdot> \<eta> \<star> g') = \<r>\<^sup>-\<^sup>1[g'] \<cdot> \<l>[g']"
proof -
have "(g' \<star> \<epsilon> \<cdot> (f \<star> inv \<phi>)) \<cdot> \<a>[g', f, g'] \<cdot> ((\<phi> \<star> f) \<cdot> \<eta> \<star> g') =
(g' \<star> \<epsilon>) \<cdot> ((g' \<star> f \<star> inv \<phi>) \<cdot> \<a>[g', f, g']) \<cdot> ((\<phi> \<star> f) \<star> g') \<cdot> (\<eta> \<star> g')"
using assms A.antipar whisker_left whisker_right comp_assoc by auto
also have "... = (g' \<star> \<epsilon>) \<cdot> (\<a>[g', f, g] \<cdot> ((g' \<star> f) \<star> inv \<phi>)) \<cdot> ((\<phi> \<star> f) \<star> g') \<cdot> (\<eta> \<star> g')"
using assms A.antipar assoc_naturality [of g' f "inv \<phi>"] by auto
also have "... = (g' \<star> \<epsilon>) \<cdot> \<a>[g', f, g] \<cdot> (((g' \<star> f) \<star> inv \<phi>) \<cdot> ((\<phi> \<star> f) \<star> g')) \<cdot> (\<eta> \<star> g')"
using comp_assoc by simp
also have "... = (g' \<star> \<epsilon>) \<cdot> (\<a>[g', f, g] \<cdot> ((\<phi> \<star> f) \<star> g)) \<cdot> ((g \<star> f) \<star> inv \<phi>) \<cdot> (\<eta> \<star> g')"
proof -
have "((g' \<star> f) \<star> inv \<phi>) \<cdot> ((\<phi> \<star> f) \<star> g') = (\<phi> \<star> f) \<star> inv \<phi>"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "g' \<star> f" "\<phi> \<star> f" "inv \<phi>" g']
by auto
also have "... = ((\<phi> \<star> f) \<star> g) \<cdot> ((g \<star> f) \<star> inv \<phi>)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "\<phi> \<star> f" "g \<star> f" g "inv \<phi>"]
by auto
finally show ?thesis
using comp_assoc by simp
qed
also have "... = ((g' \<star> \<epsilon>) \<cdot> (\<phi> \<star> f \<star> g)) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> (trg g \<star> inv \<phi>)"
proof -
have "\<a>[g', f, g] \<cdot> ((\<phi> \<star> f) \<star> g) = (\<phi> \<star> f \<star> g) \<cdot> \<a>[g, f, g]"
using assms A.antipar assoc_naturality [of \<phi> f g] by auto
moreover have "((g \<star> f) \<star> inv \<phi>) \<cdot> (\<eta> \<star> g') = (\<eta> \<star> g) \<cdot> (trg g \<star> inv \<phi>)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "g \<star> f" \<eta> "inv \<phi>" g'] interchange [of \<eta> "trg g" g "inv \<phi>"]
by auto
ultimately show ?thesis
using comp_assoc by simp
qed
also have "... = ((\<phi> \<star> src g) \<cdot> (g \<star> \<epsilon>)) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> (trg g \<star> inv \<phi>)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of g' \<phi> \<epsilon> "f \<star> g"] interchange [of \<phi> g "src g" \<epsilon>]
by (metis A.counit_simps(1) A.counit_simps(2) A.counit_simps(3) in_homE)
also have "... = (\<phi> \<star> src g) \<cdot> ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot> (trg g \<star> inv \<phi>)"
using comp_assoc by simp
also have "... = ((\<phi> \<star> src g) \<cdot> \<r>\<^sup>-\<^sup>1[g]) \<cdot> \<l>[g] \<cdot> (trg g \<star> inv \<phi>)"
using assms A.antipar A.triangle_right comp_cod_arr comp_assoc
by simp
also have "... = (\<r>\<^sup>-\<^sup>1[g'] \<cdot> \<phi>) \<cdot> inv \<phi> \<cdot> \<l>[g']"
using assms A.antipar runit'_naturality [of \<phi>] lunit_naturality [of "inv \<phi>"]
by auto
also have "... = \<r>\<^sup>-\<^sup>1[g'] \<cdot> (\<phi> \<cdot> inv \<phi>) \<cdot> \<l>[g']"
using comp_assoc by simp
also have "... = \<r>\<^sup>-\<^sup>1[g'] \<cdot> \<l>[g']"
using assms comp_cod_arr comp_arr_inv' by auto
finally show ?thesis by simp
qed
qed
lemma adjunction_preserved_by_iso_left:
assumes "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "adjunction_in_bicategory V H \<a> \<i> src trg f' g ((g \<star> \<phi>) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g))"
proof
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by auto
show "ide g" by simp
show "ide f'"
using assms(2) isomorphic_def by auto
show "\<guillemotleft>(g \<star> \<phi>) \<cdot> \<eta> : src f' \<Rightarrow> g \<star> f'\<guillemotright>"
using assms A.antipar A.unit_in_hom by force
show "\<guillemotleft>\<epsilon> \<cdot> (inv \<phi> \<star> g) : f' \<star> g \<Rightarrow> src g\<guillemotright>"
using assms A.antipar by force
show "(g \<star> \<epsilon> \<cdot> (inv \<phi> \<star> g)) \<cdot> \<a>[g, f', g] \<cdot> ((g \<star> \<phi>) \<cdot> \<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
proof -
have "(g \<star> \<epsilon> \<cdot> (inv \<phi> \<star> g)) \<cdot> \<a>[g, f', g] \<cdot> ((g \<star> \<phi>) \<cdot> \<eta> \<star> g) =
(g \<star> \<epsilon>) \<cdot> ((g \<star> inv \<phi> \<star> g) \<cdot> \<a>[g, f', g]) \<cdot> ((g \<star> \<phi>) \<star> g) \<cdot> (\<eta> \<star> g)"
using assms A.antipar whisker_left whisker_right comp_assoc by auto
also have "... = (g \<star> \<epsilon>) \<cdot> (\<a>[g, f, g] \<cdot> ((g \<star> inv \<phi>) \<star> g)) \<cdot> ((g \<star> \<phi>) \<star> g) \<cdot> (\<eta> \<star> g)"
using assms A.antipar assoc_naturality [of g "inv \<phi>" g] by auto
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (((g \<star> inv \<phi>) \<star> g) \<cdot> ((g \<star> \<phi>) \<star> g)) \<cdot> (\<eta> \<star> g)"
using comp_assoc by simp
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> ((g \<star> f) \<star> g) \<cdot> (\<eta> \<star> g)"
using assms A.antipar comp_inv_arr inv_is_inverse whisker_right
whisker_left [of g "inv \<phi>" \<phi>]
by auto
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)"
using assms A.antipar comp_cod_arr by simp
also have "... = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using A.triangle_right by simp
finally show ?thesis by simp
qed
show "(\<epsilon> \<cdot> (inv \<phi> \<star> g) \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f', g, f'] \<cdot> (f' \<star> (g \<star> \<phi>) \<cdot> \<eta>) = \<l>\<^sup>-\<^sup>1[f'] \<cdot> \<r>[f']"
proof -
have "(\<epsilon> \<cdot> (inv \<phi> \<star> g) \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f', g, f'] \<cdot> (f' \<star> (g \<star> \<phi>) \<cdot> \<eta>) =
(\<epsilon> \<star> f') \<cdot> (((inv \<phi> \<star> g) \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f', g, f']) \<cdot> (f' \<star> g \<star> \<phi>) \<cdot> (f' \<star> \<eta>)"
using assms A.antipar whisker_right whisker_left comp_assoc
by auto
also have "... = (\<epsilon> \<star> f') \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> (inv \<phi> \<star> g \<star> f')) \<cdot> (f' \<star> g \<star> \<phi>) \<cdot> (f' \<star> \<eta>)"
using assms A.antipar assoc'_naturality [of "inv \<phi>" g f'] by auto
also have "... = (\<epsilon> \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> ((inv \<phi> \<star> g \<star> f') \<cdot> (f' \<star> g \<star> \<phi>)) \<cdot> (f' \<star> \<eta>)"
using comp_assoc by simp
also have "... = (\<epsilon> \<star> f') \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> (f \<star> g \<star> \<phi>)) \<cdot> (inv \<phi> \<star> g \<star> f) \<cdot> (f' \<star> \<eta>)"
proof -
have "(inv \<phi> \<star> g \<star> f') \<cdot> (f' \<star> g \<star> \<phi>) = (f \<star> g \<star> \<phi>) \<cdot> (inv \<phi> \<star> g \<star> f)"
using assms(2-3) A.antipar comp_arr_dom comp_cod_arr
interchange [of "inv \<phi>" f' "g \<star> f'" "g \<star> \<phi>"]
interchange [of f "inv \<phi>" "g \<star> \<phi>" "g \<star> f"]
by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((\<epsilon> \<star> f') \<cdot> ((f \<star> g) \<star> \<phi>)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> (inv \<phi> \<star> src f)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> (f \<star> g \<star> \<phi>) = ((f \<star> g) \<star> \<phi>) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f]"
using assms A.antipar assoc'_naturality [of f g \<phi>] by auto
moreover have "(inv \<phi> \<star> g \<star> f) \<cdot> (f' \<star> \<eta>) = (f \<star> \<eta>) \<cdot> (inv \<phi> \<star> src f)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "inv \<phi>" f' "g \<star> f" \<eta>] interchange [of f "inv \<phi>" \<eta> "src f"]
by auto
ultimately show ?thesis
using comp_assoc by simp
qed
also have "... = ((trg f \<star> \<phi>) \<cdot> (\<epsilon> \<star> f)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> (inv \<phi> \<star> src f)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of \<epsilon> "f \<star> g" f' \<phi>] interchange [of "trg f" \<epsilon> \<phi> f]
by (metis A.counit_simps(1) A.counit_simps(2) A.counit_simps(3) in_homE)
also have "... = (trg f \<star> \<phi>) \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) \<cdot> (inv \<phi> \<star> src f)"
using comp_assoc by simp
also have "... = ((trg f \<star> \<phi>) \<cdot> \<l>\<^sup>-\<^sup>1[f]) \<cdot> \<r>[f] \<cdot> (inv \<phi> \<star> src f)"
using assms A.antipar A.triangle_left comp_cod_arr comp_assoc
by simp
also have "... = (\<l>\<^sup>-\<^sup>1[f'] \<cdot> \<phi>) \<cdot> inv \<phi> \<cdot> \<r>[f']"
using assms A.antipar lunit'_naturality runit_naturality [of "inv \<phi>"] by auto
also have "... = \<l>\<^sup>-\<^sup>1[f'] \<cdot> (\<phi> \<cdot> inv \<phi>) \<cdot> \<r>[f']"
using comp_assoc by simp
also have "... = \<l>\<^sup>-\<^sup>1[f'] \<cdot> \<r>[f']"
using assms comp_cod_arr comp_arr_inv inv_is_inverse by auto
finally show ?thesis by simp
qed
qed
lemma adjoint_pair_preserved_by_iso:
assumes "adjoint_pair f g"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
and "\<guillemotleft>\<psi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<psi>"
shows "adjoint_pair f' g'"
using assms adjoint_pair_def adjunction_preserved_by_iso_left
adjunction_preserved_by_iso_right
by metis
lemma left_adjoint_preserved_by_iso:
assumes "is_left_adjoint f"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "is_left_adjoint f'"
proof -
obtain g where g: "adjoint_pair f g"
using assms by auto
have "adjoint_pair f' g"
using assms g adjoint_pair_preserved_by_iso [of f g \<phi> f' g g]
adjoint_pair_antipar [of f g]
by auto
thus ?thesis by auto
qed
lemma right_adjoint_preserved_by_iso:
assumes "is_right_adjoint g"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "is_right_adjoint g'"
proof -
obtain f where f: "adjoint_pair f g"
using assms by auto
have "adjoint_pair f g'"
using assms f adjoint_pair_preserved_by_iso [of f g f f \<phi> g']
adjoint_pair_antipar [of f g]
by auto
thus ?thesis by auto
qed
lemma left_adjoint_preserved_by_iso':
assumes "is_left_adjoint f" and "f \<cong> f'"
shows "is_left_adjoint f'"
using assms isomorphic_def left_adjoint_preserved_by_iso by blast
lemma right_adjoint_preserved_by_iso':
assumes "is_right_adjoint g" and "g \<cong> g'"
shows "is_right_adjoint g'"
using assms isomorphic_def right_adjoint_preserved_by_iso by blast
lemma obj_self_adjunction:
assumes "obj a"
shows "adjunction_in_bicategory V H \<a> \<i> src trg a a \<l>\<^sup>-\<^sup>1[a] \<r>[a]"
proof
show 1: "ide a"
using assms by auto
show "\<guillemotleft>\<l>\<^sup>-\<^sup>1[a] : src a \<Rightarrow> a \<star> a\<guillemotright>"
using assms 1 by auto
show "\<guillemotleft>\<r>[a] : a \<star> a \<Rightarrow> src a\<guillemotright>"
using assms 1 by fastforce
show "(\<r>[a] \<star> a) \<cdot> \<a>\<^sup>-\<^sup>1[a, a, a] \<cdot> (a \<star> \<l>\<^sup>-\<^sup>1[a]) = \<l>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a]"
using assms 1 canI_unitor_1 canI_associator_1(2) canI_associator_3
whisker_can_right_1 whisker_can_left_1 can_Ide_self obj_simps
by simp
show "(a \<star> \<r>[a]) \<cdot> \<a>[a, a, a] \<cdot> (\<l>\<^sup>-\<^sup>1[a] \<star> a) = \<r>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a]"
using assms 1 canI_unitor_1 canI_associator_1(2) canI_associator_3
whisker_can_right_1 whisker_can_left_1 can_Ide_self
by simp
qed
lemma obj_is_self_adjoint:
assumes "obj a"
shows "adjoint_pair a a" and "is_left_adjoint a" and "is_right_adjoint a"
using assms obj_self_adjunction adjoint_pair_def by auto
end
subsection "Pseudofunctors and Adjunctions"
context pseudofunctor
begin
lemma preserves_adjunction:
assumes "adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
shows "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof -
interpret adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms by auto
interpret A: adjunction_data_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>F f\<close> \<open>F g\<close> \<open>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)\<close>
\<open>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)\<close>
using adjunction_data_in_bicategory_axioms preserves_adjunction_data by auto
show "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof
show "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
D.lunit' (F f) \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
proof -
have 1: "D.iso (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)))"
- using antipar C.VV.ide_char C.VV.arr_char D.iso_is_arr FF_def
+ using antipar C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.iso_is_arr FF_def
by (intro D.isos_compose D.seqI, simp_all)
have "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))) =
D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f))"
proof -
have "D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F f]) =
D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f))"
using antipar assoc_coherence by simp
moreover
have "D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F f]) =
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)))"
proof -
have "D.seq (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))) \<a>\<^sub>D[F f, F g, F f]"
using antipar by fastforce
thus ?thesis
using 1 antipar D.comp_assoc
D.inv_comp [of "\<a>\<^sub>D[F f, F g, F f]" "\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))"]
by auto
qed
ultimately show ?thesis by simp
qed
moreover have 2: "D.iso (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f))"
- using antipar D.isos_compose C.VV.ide_char C.VV.arr_char cmp_simps(4)
+ using antipar D.isos_compose C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_simps(4)
by simp
ultimately have "\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] =
D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)))"
using 1 2 antipar D.invert_side_of_triangle(2) D.inv_inv D.iso_inv_iso D.arr_inv
by metis
moreover have "D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
proof -
have "D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) =
D.inv (\<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
- using antipar D.isos_compose C.VV.arr_char cmp_simps(4)
- preserves_inv D.inv_comp C.VV.cod_char
+ using antipar D.isos_compose C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_simps(4)
+ preserves_inv D.inv_comp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = (D.inv (\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
- using antipar D.inv_comp C.VV.ide_char C.VV.arr_char cmp_simps(4)
+ using antipar D.inv_comp C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_simps(4)
by simp
also have "... = ((D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
- using antipar C.VV.ide_char C.VV.arr_char by simp
+ using antipar C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = (D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using D.comp_assoc by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
proof -
have "... = ((D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D
((F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)))"
proof -
have "D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f =
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)"
using ide_left ide_right antipar D.whisker_right unit_char(2)
by (metis A.counit_simps(1) A.ide_left D.comp_assoc)
moreover have "F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) =
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
using antipar unit_char(2) D.whisker_left by simp
ultimately show ?thesis by simp
qed
also have "... = (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
(((\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D (D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D \<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(((F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f)))) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)))"
using D.comp_assoc by simp
also have "... = (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
proof -
have "((F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f)))) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)"
proof -
have "(F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) = F f \<star>\<^sub>D F (g \<star>\<^sub>C f)"
using antipar unit_char(2) D.comp_arr_inv D.inv_is_inverse
D.whisker_left [of "F f" "\<Phi> (g, f)" "D.inv (\<Phi> (g, f))"]
by simp
moreover have "D.seq (F f \<star>\<^sub>D F (g \<star>\<^sub>C f)) (F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
using antipar by fastforce
ultimately show ?thesis
using D.comp_cod_arr by auto
qed
moreover have "((\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D (D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) =
D.inv (\<Phi> (f \<star>\<^sub>C g, f))"
using antipar D.comp_arr_inv D.inv_is_inverse D.comp_cod_arr
D.whisker_right [of "F f" "\<Phi> (f, g)" "D.inv (\<Phi> (f, g))"]
by simp
ultimately show ?thesis by simp
qed
finally show ?thesis by simp
qed
also have "... = (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D
((\<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]) \<cdot>\<^sub>D
((\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C f))) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
((D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F \<epsilon> \<star>\<^sub>D F f)) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
((F f \<star>\<^sub>D F \<eta>) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)))"
using antipar D.comp_assoc D.whisker_left D.whisker_right unit_char(2)
by simp
moreover have "F \<epsilon> \<star>\<^sub>D F f = D.inv (\<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f)"
- using antipar \<Phi>.naturality [of "(\<epsilon>, f)"] C.VV.arr_char FF_def
- D.invert_side_of_triangle(1) C.VV.dom_char C.VV.cod_char
+ using antipar \<Phi>.naturality [of "(\<epsilon>, f)"] C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
+ D.invert_side_of_triangle(1) C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "F f \<star>\<^sub>D F \<eta> = D.inv (\<Phi> (f, g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>) \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f)"
- using antipar \<Phi>.naturality [of "(f, \<eta>)"] C.VV.arr_char FF_def
- D.invert_side_of_triangle(1) C.VV.dom_char C.VV.cod_char
+ using antipar \<Phi>.naturality [of "(f, \<eta>)"] C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
+ D.invert_side_of_triangle(1) C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))) \<cdot>\<^sub>D
(F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D
(F ((f \<star>\<^sub>C g) \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C g \<star>\<^sub>C f)) \<cdot>\<^sub>D
F (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
using antipar D.comp_arr_inv' D.comp_assoc by simp
also have "... = ((D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))) \<cdot>\<^sub>D
(F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "F ((f \<star>\<^sub>C g) \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C g \<star>\<^sub>C f) = F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using antipar D.comp_arr_dom D.comp_cod_arr by simp
thus ?thesis by simp
qed
also have "... = D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "(D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f)) =
D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
proof -
have "D.iso (\<Phi> (trg\<^sub>C f, f))"
using antipar by simp
moreover have "D.iso (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
using antipar unit_char(2) by simp
moreover have "D.seq (\<Phi> (trg\<^sub>C f, f)) (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
using antipar D.iso_is_arr calculation(2)
apply (intro D.seqI D.hseqI) by auto
ultimately show ?thesis
using antipar D.inv_comp unit_char(2) by simp
qed
moreover have "F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>) =
F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>))"
using antipar by simp
ultimately show ?thesis by simp
qed
also have "... = (D.lunit' (F f) \<cdot>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D
F (C.lunit' f \<cdot>\<^sub>C \<r>\<^sub>C[f]) \<cdot>\<^sub>D
(D.inv (F \<r>\<^sub>C[f]) \<cdot>\<^sub>D \<r>\<^sub>D[F f])"
proof -
have "F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) = F (C.lunit' f \<cdot>\<^sub>C \<r>\<^sub>C[f])"
using triangle_left by simp
moreover have "D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) =
D.lunit' (F f) \<cdot>\<^sub>D F \<l>\<^sub>C[f]"
proof -
have 0: "D.iso (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
using unit_char(2)
apply (intro D.isos_compose D.seqI) by auto
show ?thesis
proof -
have 1: "D.iso (F \<l>\<^sub>C[f])"
using C.iso_lunit preserves_iso by auto
moreover have "D.iso (F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
by (metis (no_types) A.ide_left D.iso_lunit ide_left lunit_coherence)
moreover have "D.inv (D.inv (F \<l>\<^sub>C[f])) = F \<l>\<^sub>C[f]"
using 1 D.inv_inv by blast
ultimately show ?thesis
by (metis 0 D.inv_comp D.invert_side_of_triangle(2) D.iso_inv_iso
D.iso_is_arr ide_left lunit_coherence)
qed
qed
moreover have "\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) = D.inv (F \<r>\<^sub>C[f]) \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
using ide_left runit_coherence preserves_iso C.iso_runit D.invert_side_of_triangle(1)
by (metis A.ide_left D.runit_simps(1))
ultimately show ?thesis by simp
qed
also have "... = D.lunit' (F f) \<cdot>\<^sub>D
((F \<l>\<^sub>C[f] \<cdot>\<^sub>D F (C.lunit' f)) \<cdot>\<^sub>D (F \<r>\<^sub>C[f] \<cdot>\<^sub>D D.inv (F \<r>\<^sub>C[f]))) \<cdot>\<^sub>D
\<r>\<^sub>D[F f]"
using D.comp_assoc by simp
also have "... = D.lunit' (F f) \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
using D.comp_cod_arr C.iso_runit C.iso_lunit preserves_iso D.comp_arr_inv'
preserves_inv
by force
finally show ?thesis by blast
qed
show "(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g) =
D.runit' (F g) \<cdot>\<^sub>D \<l>\<^sub>D[F g]"
proof -
have "\<a>\<^sub>D[F g, F f, F g] =
D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g)"
proof -
have "D.iso (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)))"
using antipar D.iso_is_arr
apply (intro D.isos_compose, auto)
by (metis C.iso_assoc D.comp_assoc D.seqE ide_left ide_right
preserves_assoc(1) preserves_iso)
moreover have "F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g) =
\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, F g]"
using antipar assoc_coherence by simp
moreover have "D.seq (F \<a>\<^sub>C[g, f, g]) (\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g))"
- using antipar C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
+ using antipar C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def
by (intro D.seqI D.hseqI') auto
ultimately show ?thesis
using D.invert_side_of_triangle(1) D.comp_assoc by auto
qed
hence "(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g) =
(F g \<star>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D
\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
using D.comp_assoc by simp
also have "... = ((F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D
(\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D ((D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g))"
proof -
have "F g \<star>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D \<Phi> (f, g) =
(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))"
proof -
have "D.seq (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) (\<Phi> (f, g))"
using antipar D.comp_assoc by simp
thus ?thesis
using antipar D.whisker_left by simp
qed
moreover have "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g =
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
using antipar D.whisker_right by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D
(((F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D D.inv (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g))) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D
((\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) =
D.inv (F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D D.inv (\<Phi> (g, f \<star>\<^sub>C g))"
proof -
have "D.iso (\<Phi> (g, f \<star>\<^sub>C g))"
using antipar by simp
moreover have "D.iso (F g \<star>\<^sub>D \<Phi> (f, g))"
using antipar by simp
moreover have "D.seq (\<Phi> (g, f \<star>\<^sub>C g)) (F g \<star>\<^sub>D \<Phi> (f, g))"
using antipar cmp_in_hom A.ide_right D.iso_is_arr
by (intro D.seqI) auto
ultimately show ?thesis
using antipar D.inv_comp by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g)) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "((\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g) =
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "(\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) = F (g \<star>\<^sub>C f) \<star>\<^sub>D F g"
using antipar D.comp_arr_inv'
D.whisker_right [of "F g" "\<Phi> (g, f)" "D.inv (\<Phi> (g, f))"]
by simp
thus ?thesis
using antipar D.comp_cod_arr D.whisker_right by simp
qed
moreover have "((F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D D.inv (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g)) =
D.inv (\<Phi> (g, f \<star>\<^sub>C g))"
using antipar D.comp_arr_inv' D.comp_cod_arr
D.whisker_left [of "F g" "\<Phi> (f, g)" "D.inv (\<Phi> (f, g))"]
by simp
ultimately show ?thesis by simp
qed
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D
((F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D D.inv (\<Phi> (g, f \<star>\<^sub>C g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D
(\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(unit (src\<^sub>C f) \<star>\<^sub>D F g)"
using antipar D.whisker_left D.whisker_right unit_char(2) D.comp_assoc by simp
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D
(F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g)) \<cdot>\<^sub>D
\<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "(F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D D.inv (\<Phi> (g, f \<star>\<^sub>C g)) = D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (g \<star>\<^sub>C \<epsilon>)"
- using antipar C.VV.arr_char \<Phi>.naturality [of "(g, \<epsilon>)"] FF_def
- D.invert_opposite_sides_of_square C.VV.dom_char C.VV.cod_char
+ using antipar C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>.naturality [of "(g, \<epsilon>)"] FF_def
+ D.invert_opposite_sides_of_square C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g) = F (\<eta> \<star>\<^sub>C g) \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g)"
- using antipar C.VV.arr_char \<Phi>.naturality [of "(\<eta>, g)"] FF_def
- C.VV.dom_char C.VV.cod_char
+ using antipar C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<Phi>.naturality [of "(\<eta>, g)"] FF_def
+ C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D
F (C.runit' g)) \<cdot>\<^sub>D (F \<l>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (src\<^sub>C f) \<star>\<^sub>D F g))"
proof -
have "F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g) = F (C.runit' g) \<cdot>\<^sub>D F \<l>\<^sub>C[g]"
using ide_left ide_right antipar triangle_right
by (metis C.comp_in_homE C.seqI' preserves_comp triangle_in_hom(2))
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = D.runit' (F g) \<cdot>\<^sub>D \<l>\<^sub>D[F g]"
proof -
have "D.inv \<r>\<^sub>D[F g] =
(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (C.runit' g)"
proof -
have "D.runit' (F g) = D.inv (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g)))"
using runit_coherence by simp
also have
"... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (C.runit' g)"
proof -
have "D.inv (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) =
D.inv (F g \<star>\<^sub>D unit (src\<^sub>C g)) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (C.runit' g)"
proof -
have "D.iso (F \<r>\<^sub>C[g])"
using preserves_iso by simp
moreover have 1: "D.iso (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g)))"
using preserves_iso unit_char(2) D.arrI D.seqE ide_right runit_coherence
by (intro D.isos_compose D.seqI, auto)
moreover have "D.seq (F \<r>\<^sub>C[g]) (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g)))"
using ide_right A.ide_right D.runit_simps(1) runit_coherence by metis
ultimately have "D.inv (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) =
D.inv (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) \<cdot>\<^sub>D F (C.runit' g)"
using C.iso_runit preserves_inv D.inv_comp by simp
moreover have "D.inv (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) =
D.inv (F g \<star>\<^sub>D unit (src\<^sub>C g)) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
proof -
have "D.seq (\<Phi> (g, src\<^sub>C g)) (F g \<star>\<^sub>D unit (src\<^sub>C g))"
using 1 antipar preserves_iso unit_char(2) by fast
(*
* TODO: The fact that auto cannot do this step is probably what is blocking
* the whole thing from being done by auto.
*)
thus ?thesis
using 1 antipar preserves_iso unit_char(2) D.inv_comp by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using antipar unit_char(2) preserves_iso by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using antipar lunit_coherence by simp
qed
finally show ?thesis by simp
qed
qed
qed
lemma preserves_adjoint_pair:
assumes "C.adjoint_pair f g"
shows "D.adjoint_pair (F f) (F g)"
using assms C.adjoint_pair_def D.adjoint_pair_def preserves_adjunction by blast
lemma preserves_left_adjoint:
assumes "C.is_left_adjoint f"
shows "D.is_left_adjoint (F f)"
using assms preserves_adjoint_pair by auto
lemma preserves_right_adjoint:
assumes "C.is_right_adjoint g"
shows "D.is_right_adjoint (F g)"
using assms preserves_adjoint_pair by auto
end
context equivalence_pseudofunctor
begin
lemma reflects_adjunction:
assumes "C.ide f" and "C.ide g"
and "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>" and "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>"
and "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
shows "adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
proof -
let ?\<eta>' = "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)"
let ?\<epsilon>' = "D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)"
interpret A': adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close> ?\<eta>' ?\<epsilon>'
using assms(5) by auto
interpret A: adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms(1-4) by (unfold_locales, auto)
show ?thesis
proof
show "(\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>) = \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f]"
proof -
have 1: "C.par ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f])"
using assms A.antipar by simp
moreover have "F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) = F (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f])"
proof -
have "F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) =
F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>)"
using 1 by (metis C.seqE preserves_comp)
also have "... =
(F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
(\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>))"
using assms A.antipar preserves_assoc(2) D.comp_assoc by auto
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D ((F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
((F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta>)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (F \<epsilon> \<star>\<^sub>D F f)"
- using assms \<Phi>.naturality [of "(\<epsilon>, f)"] FF_def C.VV.arr_char
- C.VV.dom_char C.VV.cod_char
+ using assms \<Phi>.naturality [of "(\<epsilon>, f)"] FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "D.inv (\<Phi> (f, g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>) =
(F f \<star>\<^sub>D F \<eta>) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F (f \<star>\<^sub>C \<eta>) \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) = \<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta>)"
- using assms \<Phi>.naturality [of "(f, \<eta>)"] FF_def C.VV.arr_char A.antipar
- C.VV.dom_char C.VV.cod_char
+ using assms \<Phi>.naturality [of "(f, \<eta>)"] FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C A.antipar
+ C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
thus ?thesis
- using assms A.antipar cmp_components_are_iso C.VV.arr_char cmp_in_hom
+ using assms A.antipar cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_in_hom
FF_def C.VV.dom_simp C.VV.cod_simp
D.invert_opposite_sides_of_square
[of "\<Phi> (f, g \<star>\<^sub>C f)" "F f \<star>\<^sub>D F \<eta>" "F (f \<star>\<^sub>C \<eta>)" "\<Phi> (f, src\<^sub>C f)"]
by fastforce
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta>) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using assms A.antipar cmp_in_hom A.ide_left A.ide_right A'.ide_left A'.ide_right
D.whisker_left [of "F f" "D.inv (\<Phi> (g, f))" "F \<eta>"]
D.whisker_right [of "F f" "F \<epsilon>" "\<Phi> (f, g)"]
by (metis A'.counit_in_vhom A'.unit_simps(1)D.arrI D.comp_assoc
D.src.preserves_reflects_arr D.src_vcomp D.vseq_implies_hpar(1) cmp_simps(2))
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>' \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D ?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) = unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>'"
proof -
have "D.iso (unit (trg\<^sub>C f))"
using A.ide_left C.ideD(1) unit_char(2) by blast
thus ?thesis
by (metis A'.counit_simps(1) D.comp_assoc D.comp_cod_arr D.inv_is_inverse
D.seqE D.comp_arr_inv)
qed
moreover have "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> = ?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
using assms(2) unit_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr
by (metis A'.unit_simps(1) A.antipar(1) C.ideD(1) C.obj_trg
D.invert_side_of_triangle(2))
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D ((unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(?\<epsilon>' \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D ((F f \<star>\<^sub>D ?\<eta>') \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f)))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using assms A.antipar A'.antipar unit_char D.whisker_left D.whisker_right
by simp
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D
((?\<epsilon>' \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D ?\<eta>')) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using D.comp_assoc by simp
also have "... = (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]) \<cdot>\<^sub>D
\<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using A'.triangle_left D.comp_assoc by simp
also have "... = F \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>D F \<r>\<^sub>C[f]"
using assms A.antipar preserves_lunit(2) preserves_runit(1) by simp
also have "... = F (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f])"
using assms by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using is_faithful by blast
qed
show "(g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g) = \<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g]"
proof -
have 1: "C.par ((g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C g f g \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g)) (\<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g])"
using assms A.antipar by auto
moreover have "F ((g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g)) = F (\<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g])"
proof -
have "F ((g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C g f g \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g)) =
F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g)"
using 1 by auto
also have "... = (F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D \<Phi> (g, f \<star>\<^sub>C g)) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g \<star>\<^sub>C f, g)) \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g))"
using assms A.antipar preserves_assoc(1) [of g f g] D.comp_assoc by auto
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D ((F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
((D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g)) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D \<Phi> (g, f \<star>\<^sub>C g) = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<epsilon>)"
- using assms \<Phi>.naturality [of "(g, \<epsilon>)"] FF_def C.VV.arr_char
+ using assms \<Phi>.naturality [of "(g, \<epsilon>)"] FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
C.VV.dom_simp C.VV.cod_simp
by auto
moreover have "D.inv (\<Phi> (g \<star>\<^sub>C f, g)) \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g) =
(F \<eta> \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "F (\<eta> \<star>\<^sub>C g) \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g) = \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g)"
- using assms \<Phi>.naturality [of "(\<eta>, g)"] FF_def C.VV.arr_char A.antipar
+ using assms \<Phi>.naturality [of "(\<eta>, g)"] FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C A.antipar
C.VV.dom_simp C.VV.cod_simp
by auto
thus ?thesis
- using assms A.antipar cmp_components_are_iso C.VV.arr_char FF_def
+ using assms A.antipar cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def
C.VV.dom_simp C.VV.cod_simp
D.invert_opposite_sides_of_square
[of "\<Phi> (g \<star>\<^sub>C f, g)" "F \<eta> \<star>\<^sub>D F g" "F (\<eta> \<star>\<^sub>C g)" "\<Phi> (trg\<^sub>C g, g)"]
by fastforce
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have " ... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "(F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)) = F g \<star>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)"
using assms A.antipar D.whisker_left
by (metis A'.counit_simps(1) A'.ide_right D.seqE)
moreover have "(D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g) =
D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<star>\<^sub>D F g"
using assms A.antipar D.whisker_right by simp
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) = unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>'"
using unit_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr
by (metis A'.counit_simps(1) C.ideD(1) C.obj_trg D.seqE assms(1))
moreover have "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> = ?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
using unit_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr
by (metis A'.unit_simps(1) A.unit_simps(1) A.unit_simps(5)
C.obj_trg D.invert_side_of_triangle(2))
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (trg\<^sub>C f)) \<cdot>\<^sub>D
((F g \<star>\<^sub>D ?\<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D (?\<eta>' \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(D.inv (unit (src\<^sub>C f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
using assms A.antipar unit_char D.whisker_left D.whisker_right D.comp_assoc
by simp
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (trg\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<cdot>\<^sub>D
\<l>\<^sub>D[F g] \<cdot>\<^sub>D (D.inv (unit (src\<^sub>C f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
using A'.triangle_right D.comp_assoc by simp
also have "... = F \<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>D F \<l>\<^sub>C[g]"
using assms A.antipar preserves_lunit(1) preserves_runit(2) D.comp_assoc
by simp
also have "... = F (\<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g])"
using assms by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using is_faithful by blast
qed
qed
qed
lemma reflects_adjoint_pair:
assumes "C.ide f" and "C.ide g"
and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C f"
and "D.adjoint_pair (F f) (F g)"
shows "C.adjoint_pair f g"
proof -
obtain \<eta>' \<epsilon>' where A': "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) \<eta>' \<epsilon>'"
using assms D.adjoint_pair_def by auto
interpret A': adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close> \<eta>' \<epsilon>'
using A' by auto
have 1: "\<guillemotleft>\<Phi> (g, f) \<cdot>\<^sub>D \<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f)) : F (src\<^sub>C f) \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C f)\<guillemotright>"
using assms unit_char [of "src\<^sub>C f"] A'.unit_in_hom
by (intro D.comp_in_homI, auto)
have 2: "\<guillemotleft>unit (trg\<^sub>C f) \<cdot>\<^sub>D \<epsilon>' \<cdot>\<^sub>D D.inv (\<Phi> (f, g)): F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
using assms cmp_in_hom [of f g] unit_char [of "trg\<^sub>C f"] A'.counit_in_hom
by (intro D.comp_in_homI, auto)
obtain \<eta> where \<eta>: "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright> \<and>
F \<eta> = \<Phi> (g, f) \<cdot>\<^sub>D \<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
using assms 1 A'.unit_in_hom cmp_in_hom locally_full by fastforce
have \<eta>': "\<eta>' = D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)"
using assms 1 \<eta> cmp_in_hom D.iso_inv_iso cmp_components_are_iso unit_char(2)
D.invert_side_of_triangle(1) [of "F \<eta>" "\<Phi> (g, f)" "\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"]
D.invert_side_of_triangle(2) [of "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta>" \<eta>' "D.inv (unit (src\<^sub>C f))"]
by (metis (no_types, lifting) C.ideD(1) C.obj_trg D.arrI D.comp_assoc D.inv_inv)
obtain \<epsilon> where \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C trg\<^sub>C f\<guillemotright> \<and>
F \<epsilon> = unit (trg\<^sub>C f) \<cdot>\<^sub>D \<epsilon>' \<cdot>\<^sub>D D.inv (\<Phi> (f, g))"
using assms 2 A'.counit_in_hom cmp_in_hom locally_full by fastforce
have \<epsilon>': "\<epsilon>' = D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)"
using assms 2 \<epsilon> cmp_in_hom D.iso_inv_iso unit_char(2) D.comp_assoc
D.invert_side_of_triangle(1) [of "F \<epsilon>" "unit (trg\<^sub>C f)" "\<epsilon>' \<cdot>\<^sub>D D.inv (\<Phi> (f, g))"]
D.invert_side_of_triangle(2) [of "D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>" \<epsilon>' "D.inv (\<Phi> (f, g))"]
by (metis (no_types, lifting) C.arrI C.ideD(1) C.obj_trg D.inv_inv cmp_components_are_iso
preserves_arr)
have "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
using A'.adjunction_in_bicategory_axioms \<eta>' \<epsilon>' by simp
hence "adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
using assms \<eta> \<epsilon> reflects_adjunction by simp
thus ?thesis
using C.adjoint_pair_def by auto
qed
lemma reflects_left_adjoint:
assumes "C.ide f" and "D.is_left_adjoint (F f)"
shows "C.is_left_adjoint f"
proof -
obtain g' where g': "D.adjoint_pair (F f) g'"
using assms D.adjoint_pair_def by auto
obtain g where g: "\<guillemotleft>g : trg\<^sub>C f \<rightarrow>\<^sub>C src\<^sub>C f\<guillemotright> \<and> C.ide g \<and> D.isomorphic (F g) g'"
using assms g' locally_essentially_surjective [of "trg\<^sub>C f" "src\<^sub>C f" g']
D.adjoint_pair_antipar [of "F f" g']
by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : g' \<Rightarrow>\<^sub>D F g\<guillemotright> \<and> D.iso \<phi>"
using g D.isomorphic_def D.isomorphic_symmetric by metis
have "D.adjoint_pair (F f) (F g)"
using assms g g' \<phi> D.adjoint_pair_preserved_by_iso [of "F f" g' "F f" "F f" \<phi> "F g"]
by auto
thus ?thesis
using assms g reflects_adjoint_pair [of f g] D.adjoint_pair_antipar C.in_hhom_def
by auto
qed
lemma reflects_right_adjoint:
assumes "C.ide g" and "D.is_right_adjoint (F g)"
shows "C.is_right_adjoint g"
proof -
obtain f' where f': "D.adjoint_pair f' (F g)"
using assms D.adjoint_pair_def by auto
obtain f where f: "\<guillemotleft>f : trg\<^sub>C g \<rightarrow>\<^sub>C src\<^sub>C g\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) f'"
using assms f' locally_essentially_surjective [of "trg\<^sub>C g" "src\<^sub>C g" f']
D.adjoint_pair_antipar [of f' "F g"]
by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : f' \<Rightarrow>\<^sub>D F f\<guillemotright> \<and> D.iso \<phi>"
using f D.isomorphic_def D.isomorphic_symmetric by metis
have "D.adjoint_pair (F f) (F g)"
using assms f f' \<phi> D.adjoint_pair_preserved_by_iso [of f' "F g" \<phi> "F f" "F g" "F g"]
by auto
thus ?thesis
using assms f reflects_adjoint_pair [of f g] D.adjoint_pair_antipar C.in_hhom_def
by auto
qed
end
subsection "Composition of Adjunctions"
text \<open>
We first consider the strict case, then extend to all bicategories using strictification.
\<close>
locale composite_adjunction_in_strict_bicategory =
strict_bicategory V H \<a> \<i> src trg +
fg: adjunction_in_strict_bicategory V H \<a> \<i> src trg f g \<zeta> \<xi> +
hk: adjunction_in_strict_bicategory V H \<a> \<i> src trg h k \<sigma> \<tau>
for V :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and f :: "'a"
and g :: "'a"
and \<zeta> :: "'a"
and \<xi> :: "'a"
and h :: "'a"
and k :: "'a"
and \<sigma> :: "'a"
and \<tau> :: "'a" +
assumes composable: "src h = trg f"
begin
abbreviation \<eta>
where "\<eta> \<equiv> (g \<star> \<sigma> \<star> f) \<cdot> \<zeta>"
abbreviation \<epsilon>
where "\<epsilon> \<equiv> \<tau> \<cdot> (h \<star> \<xi> \<star> k)"
interpretation adjunction_data_in_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
proof
show "ide (h \<star> f)"
using composable by simp
show "ide (g \<star> k)"
using fg.antipar hk.antipar composable by simp
show "\<guillemotleft>\<eta> : src (h \<star> f) \<Rightarrow> (g \<star> k) \<star> h \<star> f\<guillemotright>"
proof
show "\<guillemotleft>\<zeta> : src (h \<star> f) \<Rightarrow> g \<star> f\<guillemotright>"
using fg.antipar hk.antipar composable \<open>ide (h \<star> f)\<close> by auto
show "\<guillemotleft>g \<star> \<sigma> \<star> f : g \<star> f \<Rightarrow> (g \<star> k) \<star> h \<star> f\<guillemotright>"
proof -
have "\<guillemotleft>g \<star> \<sigma> \<star> f : g \<star> trg f \<star> f \<Rightarrow> g \<star> (k \<star> h) \<star> f\<guillemotright>"
using fg.antipar hk.antipar composable hk.unit_in_hom
apply (intro hcomp_in_vhom) by auto
thus ?thesis
using hcomp_obj_arr hcomp_assoc by fastforce
qed
qed
show "\<guillemotleft>\<epsilon> : (h \<star> f) \<star> g \<star> k \<Rightarrow> src (g \<star> k)\<guillemotright>"
proof
show "\<guillemotleft>h \<star> \<xi> \<star> k : (h \<star> f) \<star> g \<star> k \<Rightarrow> h \<star> k\<guillemotright>"
proof -
have "\<guillemotleft>h \<star> \<xi> \<star> k : h \<star> (f \<star> g) \<star> k \<Rightarrow> h \<star> trg f \<star> k\<guillemotright>"
using composable fg.antipar(1-2) hk.antipar(1) by fastforce
thus ?thesis
using fg.antipar hk.antipar composable hk.unit_in_hom hcomp_obj_arr hcomp_assoc
by simp
qed
show "\<guillemotleft>\<tau> : h \<star> k \<Rightarrow> src (g \<star> k)\<guillemotright>"
using fg.antipar hk.antipar composable hk.unit_in_hom by auto
qed
qed
sublocale adjunction_in_strict_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
proof
show "(\<epsilon> \<star> h \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[h \<star> f, g \<star> k, h \<star> f] \<cdot> ((h \<star> f) \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[h \<star> f] \<cdot> \<r>[h \<star> f]"
proof -
have "(\<epsilon> \<star> h \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[h \<star> f, g \<star> k, h \<star> f] \<cdot> ((h \<star> f) \<star> \<eta>) =
(\<tau> \<cdot> (h \<star> \<xi> \<star> k) \<star> h \<star> f) \<cdot> ((h \<star> f) \<star> (g \<star> \<sigma> \<star> f) \<cdot> \<zeta>)"
using fg.antipar hk.antipar composable strict_assoc comp_ide_arr
ide_left ide_right antipar(1) antipar(2)
by (metis arrI seqE strict_assoc' triangle_in_hom(1))
also have "... = (\<tau> \<star> h \<star> f) \<cdot> ((h \<star> \<xi> \<star> (k \<star> h) \<star> f) \<cdot> (h \<star> (f \<star> g) \<star> \<sigma> \<star> f)) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable whisker_left [of "h \<star> f"] whisker_right
comp_assoc hcomp_assoc
by simp
also have "... = (\<tau> \<star> h \<star> f) \<cdot> (h \<star> (\<xi> \<star> (k \<star> h)) \<cdot> ((f \<star> g) \<star> \<sigma>) \<star> f) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = (\<tau> \<star> h \<star> f) \<cdot> (h \<star> (trg f \<star> \<sigma>) \<cdot> (\<xi> \<star> trg f) \<star> f) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable comp_arr_dom comp_cod_arr
interchange [of \<xi> "f \<star> g" "k \<star> h" \<sigma>] interchange [of "trg f" \<xi> \<sigma> "trg f"]
by (metis fg.counit_simps(1) fg.counit_simps(2) fg.counit_simps(3)
hk.unit_simps(1) hk.unit_simps(2) hk.unit_simps(3))
also have "... = (\<tau> \<star> h \<star> f) \<cdot> (h \<star> \<sigma> \<cdot> \<xi> \<star> f) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable hcomp_obj_arr hcomp_arr_obj
by (metis fg.counit_simps(1) fg.counit_simps(4) hk.unit_simps(1) hk.unit_simps(5)
obj_src)
also have "... = ((\<tau> \<star> h \<star> f) \<cdot> (h \<star> \<sigma> \<star> f)) \<cdot> ((h \<star> \<xi> \<star> f) \<cdot> (h \<star> f \<star> \<zeta>))"
using fg.antipar hk.antipar composable whisker_left whisker_right comp_assoc
by simp
also have "... = ((\<tau> \<star> h) \<cdot> (h \<star> \<sigma>) \<star> f) \<cdot> (h \<star> (\<xi> \<star> f) \<cdot> (f \<star> \<zeta>))"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = h \<star> f"
using fg.antipar hk.antipar composable fg.triangle_left hk.triangle_left
by simp
also have "... = \<l>\<^sup>-\<^sup>1[h \<star> f] \<cdot> \<r>[h \<star> f]"
using fg.antipar hk.antipar composable strict_lunit' strict_runit by simp
finally show ?thesis by simp
qed
show "((g \<star> k) \<star> \<epsilon>) \<cdot> \<a>[g \<star> k, h \<star> f, g \<star> k] \<cdot> (\<eta> \<star> g \<star> k) = \<r>\<^sup>-\<^sup>1[g \<star> k] \<cdot> \<l>[g \<star> k]"
proof -
have "((g \<star> k) \<star> \<epsilon>) \<cdot> \<a>[g \<star> k, h \<star> f, g \<star> k] \<cdot> (\<eta> \<star> g \<star> k) =
((g \<star> k) \<star> \<tau> \<cdot> (h \<star> \<xi> \<star> k)) \<cdot> ((g \<star> \<sigma> \<star> f) \<cdot> \<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable strict_assoc comp_ide_arr
ide_left ide_right
by (metis antipar(1) antipar(2) arrI seqE triangle_in_hom(2))
also have "... = (g \<star> k \<star> \<tau>) \<cdot> ((g \<star> (k \<star> h) \<star> \<xi> \<star> k) \<cdot> (g \<star> \<sigma> \<star> (f \<star> g) \<star> k)) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable whisker_left [of "g \<star> k"] whisker_right
comp_assoc hcomp_assoc
by simp
also have "... = (g \<star> k \<star> \<tau>) \<cdot> (g \<star> ((k \<star> h) \<star> \<xi>) \<cdot> (\<sigma> \<star> f \<star> g) \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = (g \<star> k \<star> \<tau>) \<cdot> (g \<star> (\<sigma> \<star> src g) \<cdot> (src g \<star> \<xi>) \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable interchange [of "k \<star> h" \<sigma> \<xi> "f \<star> g"]
interchange [of \<sigma> "src g" "src g" \<xi>] comp_arr_dom comp_cod_arr
by (metis fg.counit_simps(1) fg.counit_simps(2) fg.counit_simps(3)
hk.unit_simps(1) hk.unit_simps(2) hk.unit_simps(3))
also have "... = (g \<star> k \<star> \<tau>) \<cdot> (g \<star> \<sigma> \<cdot> \<xi> \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable hcomp_obj_arr [of "src g" \<xi>]
hcomp_arr_obj [of \<sigma> "src g"]
by simp
also have "... = ((g \<star> k \<star> \<tau>) \<cdot> (g \<star> \<sigma> \<star> k)) \<cdot> (g \<star> \<xi> \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable whisker_left whisker_right comp_assoc
by simp
also have "... = (g \<star> (k \<star> \<tau>) \<cdot> (\<sigma> \<star> k)) \<cdot> ((g \<star> \<xi>) \<cdot> (\<zeta> \<star> g) \<star> k)"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = g \<star> k"
using fg.antipar hk.antipar composable fg.triangle_right hk.triangle_right
by simp
also have "... = \<r>\<^sup>-\<^sup>1[g \<star> k] \<cdot> \<l>[g \<star> k]"
using fg.antipar hk.antipar composable strict_lunit strict_runit' by simp
finally show ?thesis by simp
qed
qed
lemma is_adjunction_in_strict_bicategory:
shows "adjunction_in_strict_bicategory V H \<a> \<i> src trg (h \<star> f) (g \<star> k) \<eta> \<epsilon>"
..
end
context strict_bicategory
begin
lemma left_adjoints_compose:
assumes "is_left_adjoint f" and "is_left_adjoint f'" and "src f' = trg f"
shows "is_left_adjoint (f' \<star> f)"
proof -
obtain g \<eta> \<epsilon> where fg: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret fg: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by auto
obtain g' \<eta>' \<epsilon>' where f'g': "adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret f'g': adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'
using f'g' by auto
interpret f'fgg': composite_adjunction_in_strict_bicategory V H \<a> \<i> src trg
f g \<eta> \<epsilon> f' g' \<eta>' \<epsilon>'
using assms apply unfold_locales by simp
have "adjoint_pair (f' \<star> f) (g \<star> g')"
using adjoint_pair_def f'fgg'.adjunction_in_bicategory_axioms by auto
thus ?thesis by auto
qed
lemma right_adjoints_compose:
assumes "is_right_adjoint g" and "is_right_adjoint g'" and "src g = trg g'"
shows "is_right_adjoint (g \<star> g')"
proof -
obtain f \<eta> \<epsilon> where fg: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret fg: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by auto
obtain f' \<eta>' \<epsilon>' where f'g': "adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret f'g': adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'
using f'g' by auto
interpret f'fgg': composite_adjunction_in_strict_bicategory V H \<a> \<i> src trg
f g \<eta> \<epsilon> f' g' \<eta>' \<epsilon>'
using assms fg.antipar f'g'.antipar apply unfold_locales by simp
have "adjoint_pair (f' \<star> f) (g \<star> g')"
using adjoint_pair_def f'fgg'.adjunction_in_bicategory_axioms by auto
thus ?thesis by auto
qed
end
text \<open>
We now use strictification to extend the preceding results to an arbitrary bicategory.
We only prove that ``left adjoints compose'' and ``right adjoints compose'';
I did not work out formulas for the unit and counit of the composite adjunction in the
non-strict case.
\<close>
context bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: fully_faithful_functor V S.vcomp S.UP
using S.UP_is_fully_faithful_functor by auto
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
using S.UP_is_equivalence_pseudofunctor by auto
lemma left_adjoints_compose:
assumes "is_left_adjoint f" and "is_left_adjoint f'" and "src f = trg f'"
shows "is_left_adjoint (f \<star> f')"
proof -
have "S.is_left_adjoint (S.UP f) \<and> S.is_left_adjoint (S.UP f')"
using assms UP.preserves_left_adjoint by simp
moreover have "S.src (S.UP f) = S.trg (S.UP f')"
using assms left_adjoint_is_ide by simp
ultimately have "S.is_left_adjoint (S.hcomp (S.UP f) (S.UP f'))"
using S.left_adjoints_compose by simp
moreover have "S.isomorphic (S.hcomp (S.UP f) (S.UP f')) (S.UP (f \<star> f'))"
proof -
have "\<guillemotleft>S.cmp\<^sub>U\<^sub>P (f, f') : S.hcomp (S.UP f) (S.UP f') \<Rightarrow>\<^sub>S S.UP (f \<star> f')\<guillemotright>"
using assms left_adjoint_is_ide UP.cmp_in_hom by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (f, f'))"
using assms left_adjoint_is_ide by simp
ultimately show ?thesis
using S.isomorphic_def by blast
qed
ultimately have "S.is_left_adjoint (S.UP (f \<star> f'))"
using S.left_adjoint_preserved_by_iso S.isomorphic_def by blast
thus "is_left_adjoint (f \<star> f')"
using assms left_adjoint_is_ide UP.reflects_left_adjoint by simp
qed
lemma right_adjoints_compose:
assumes "is_right_adjoint g" and "is_right_adjoint g'" and "src g' = trg g"
shows "is_right_adjoint (g' \<star> g)"
proof -
have "S.is_right_adjoint (S.UP g) \<and> S.is_right_adjoint (S.UP g')"
using assms UP.preserves_right_adjoint by simp
moreover have "S.src (S.UP g') = S.trg (S.UP g)"
using assms right_adjoint_is_ide by simp
ultimately have "S.is_right_adjoint (S.hcomp (S.UP g') (S.UP g))"
using S.right_adjoints_compose by simp
moreover have "S.isomorphic (S.hcomp (S.UP g') (S.UP g)) (S.UP (g' \<star> g))"
proof -
have "\<guillemotleft>S.cmp\<^sub>U\<^sub>P (g', g) : S.hcomp (S.UP g') (S.UP g) \<Rightarrow>\<^sub>S S.UP (g' \<star> g)\<guillemotright>"
using assms right_adjoint_is_ide UP.cmp_in_hom by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (g', g))"
using assms right_adjoint_is_ide by simp
ultimately show ?thesis
using S.isomorphic_def by blast
qed
ultimately have "S.is_right_adjoint (S.UP (g' \<star> g))"
using S.right_adjoint_preserved_by_iso S.isomorphic_def by blast
thus "is_right_adjoint (g' \<star> g)"
using assms right_adjoint_is_ide UP.reflects_right_adjoint by simp
qed
end
subsection "Choosing Right Adjoints"
text \<open>
It will be useful in various situations to suppose that we have made a choice of
right adjoint for each left adjoint ({\it i.e.} each ``map'') in a bicategory.
\<close>
locale chosen_right_adjoints =
bicategory
begin
(* Global notation is evil! *)
no_notation Transitive_Closure.rtrancl ("(_\<^sup>*)" [1000] 999)
definition some_right_adjoint ("_\<^sup>*" [1000] 1000)
where "f\<^sup>* \<equiv> SOME g. adjoint_pair f g"
definition some_unit
where "some_unit f \<equiv> SOME \<eta>. \<exists>\<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* \<eta> \<epsilon>"
definition some_counit
where "some_counit f \<equiv>
SOME \<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* (some_unit f) \<epsilon>"
lemma left_adjoint_extends_to_adjunction:
assumes "is_left_adjoint f"
shows "adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* (some_unit f) (some_counit f)"
using assms some_right_adjoint_def adjoint_pair_def some_unit_def some_counit_def
someI_ex [of "\<lambda>g. adjoint_pair f g"]
someI_ex [of "\<lambda>\<eta>. \<exists>\<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* (some_unit f) \<epsilon>"]
by auto
lemma left_adjoint_extends_to_adjoint_pair:
assumes "is_left_adjoint f"
shows "adjoint_pair f f\<^sup>*"
using assms adjoint_pair_def left_adjoint_extends_to_adjunction by blast
lemma right_adjoint_in_hom [intro]:
assumes "is_left_adjoint f"
shows "\<guillemotleft>f\<^sup>* : trg f \<rightarrow> src f\<guillemotright>"
and "\<guillemotleft>f\<^sup>* : f\<^sup>* \<Rightarrow> f\<^sup>*\<guillemotright>"
using assms left_adjoint_extends_to_adjoint_pair adjoint_pair_antipar [of f "f\<^sup>*"]
by auto
lemma right_adjoint_simps [simp]:
assumes "is_left_adjoint f"
shows "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 right_adjoint_in_hom left_adjoint_extends_to_adjoint_pair apply auto
using assms right_adjoint_is_ide [of "f\<^sup>*"] by blast
end
locale map_in_bicategory =
bicategory + chosen_right_adjoints +
fixes f :: 'a
assumes is_map: "is_left_adjoint f"
begin
abbreviation \<eta>
where "\<eta> \<equiv> some_unit f"
abbreviation \<epsilon>
where "\<epsilon> \<equiv> some_counit f"
sublocale adjunction_in_bicategory V H \<a> \<i> src trg f \<open>f\<^sup>*\<close> \<eta> \<epsilon>
using is_map left_adjoint_extends_to_adjunction by simp
end
subsection "Equivalences Refine to Adjoint Equivalences"
text \<open>
In this section, we show that, just as an equivalence between categories can always
be refined to an adjoint equivalence, an internal equivalence in a bicategory can also
always be so refined.
The proof, which follows that of Theorem 3.3 from \cite{nlab-adjoint-equivalence},
makes use of the fact that if an internal equivalence satisfies one of the triangle
identities, then it also satisfies the other.
\<close>
locale adjoint_equivalence_in_bicategory =
equivalence_in_bicategory +
adjunction_in_bicategory
begin
lemma dual_adjoint_equivalence:
shows "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg g f (inv \<epsilon>) (inv \<eta>)"
proof -
interpret gf: equivalence_in_bicategory V H \<a> \<i> src trg g f \<open>inv \<epsilon>\<close> \<open>inv \<eta>\<close>
using dual_equivalence by simp
show ?thesis
proof
show "(inv \<eta> \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[g, f, g] \<cdot> (g \<star> inv \<epsilon>) = \<l>\<^sup>-\<^sup>1[g] \<cdot> \<r>[g]"
proof -
have "(inv \<eta> \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[g, f, g] \<cdot> (g \<star> inv \<epsilon>) =
inv ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g))"
using antipar inv_comp isos_compose comp_assoc by simp
also have "... = inv (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g])"
using triangle_right by simp
also have "... = \<l>\<^sup>-\<^sup>1[g] \<cdot> \<r>[g]"
using inv_comp by simp
finally show ?thesis
by blast
qed
show "(f \<star> inv \<eta>) \<cdot> \<a>[f, g, f] \<cdot> (inv \<epsilon> \<star> f) = \<r>\<^sup>-\<^sup>1[f] \<cdot> \<l>[f]"
proof -
have "(f \<star> inv \<eta>) \<cdot> \<a>[f, g, f] \<cdot> (inv \<epsilon> \<star> f) =
inv ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>))"
using antipar inv_comp isos_compose comp_assoc by simp
also have "... = inv (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
using triangle_left by simp
also have "... = \<r>\<^sup>-\<^sup>1[f] \<cdot> \<l>[f]"
using inv_comp by simp
finally show ?thesis by blast
qed
qed
qed
end
context bicategory
begin
lemma adjoint_equivalence_preserved_by_iso_right:
assumes "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g' ((\<phi> \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (f \<star> inv \<phi>))"
proof -
interpret fg: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by simp
interpret fg': adjunction_in_bicategory V H \<a> \<i> src trg f g' \<open>(\<phi> \<star> f) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (f \<star> inv \<phi>)\<close>
using assms fg.adjunction_in_bicategory_axioms adjunction_preserved_by_iso_right
by simp
interpret fg': equivalence_in_bicategory V H \<a> \<i> src trg f g' \<open>(\<phi> \<star> f) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (f \<star> inv \<phi>)\<close>
using assms fg.equivalence_in_bicategory_axioms equivalence_preserved_by_iso_right
by simp
show ?thesis ..
qed
lemma adjoint_equivalence_preserved_by_iso_left:
assumes "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f' g ((g \<star> \<phi>) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g))"
proof -
interpret fg: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by simp
interpret fg': adjunction_in_bicategory V H \<a> \<i> src trg f' g \<open>(g \<star> \<phi>) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (inv \<phi> \<star> g)\<close>
using assms fg.adjunction_in_bicategory_axioms adjunction_preserved_by_iso_left
by simp
interpret fg': equivalence_in_bicategory V H \<a> \<i> src trg f' g \<open>(g \<star> \<phi>) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (inv \<phi> \<star> g)\<close>
using assms fg.equivalence_in_bicategory_axioms equivalence_preserved_by_iso_left
by simp
show ?thesis ..
qed
end
context strict_bicategory
begin
notation isomorphic (infix "\<cong>" 50)
lemma equivalence_refines_to_adjoint_equivalence:
assumes "equivalence_map f" and "\<guillemotleft>g : trg f \<rightarrow> src f\<guillemotright>" and "ide g"
and "\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>" and "iso \<eta>"
shows "\<exists>!\<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
obtain g' \<eta>' \<epsilon>' where E': "equivalence_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'"
using assms equivalence_map_def by auto
interpret E': equivalence_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'
using E' by auto
let ?a = "src f" and ?b = "trg f"
(* TODO: in_homE cannot be applied automatically to a conjunction. Must break down! *)
have f_in_hhom: "\<guillemotleft>f : ?a \<rightarrow> ?b\<guillemotright>" and ide_f: "ide f"
using assms equivalence_map_def by auto
have g_in_hhom: "\<guillemotleft>g : ?b \<rightarrow> ?a\<guillemotright>" and ide_g: "ide g"
using assms by auto
have g'_in_hhom: "\<guillemotleft>g' : ?b \<rightarrow> ?a\<guillemotright>" and ide_g': "ide g'"
using assms f_in_hhom E'.antipar by auto
have \<eta>_in_hom: "\<guillemotleft>\<eta> : ?a \<Rightarrow> g \<star> f\<guillemotright>" and iso_\<eta>: "iso \<eta>"
using assms by auto
have a: "obj ?a" and b: "obj ?b"
using f_in_hhom by auto
have \<eta>_in_hhom: "\<guillemotleft>\<eta> : ?a \<rightarrow> ?a\<guillemotright>"
using a \<eta>_in_hom
by (metis arrI in_hhom_def obj_simps(2-3) vconn_implies_hpar(1-2))
text \<open>
The following is quoted from \cite{nlab-adjoint-equivalence}:
\begin{quotation}
``Since \<open>g \<cong> gfg' \<cong> g'\<close>, the isomorphism \<open>fg' \<cong> 1\<close> also induces an isomorphism \<open>fg \<cong> 1\<close>,
which we denote \<open>\<xi>\<close>. Now \<open>\<eta>\<close> and \<open>\<xi>\<close> may not satisfy the zigzag identities, but if we
define \<open>\<epsilon>\<close> by \<open>\<xi> \<cdot> (f \<star> \<eta>\<^sup>-\<^sup>1) \<cdot> (f \<star> g \<star> \<xi>\<^sup>-\<^sup>1) : f \<star> g \<Rightarrow> 1\<close>, then we can verify,
using string diagram notation as above,
that \<open>\<epsilon>\<close> satisfies one zigzag identity, and hence (by the previous lemma) also the other.
Finally, if \<open>\<epsilon>': fg \<Rightarrow> 1\<close> is any other isomorphism satisfying the zigzag identities
with \<open>\<eta>\<close>, then we have:
\[
\<open>\<epsilon>' = \<epsilon>' \<cdot> (\<epsilon> f g) \<cdot> (f \<eta> g) = \<epsilon> \<cdot> (f g \<epsilon>') \<cdot> (f \<eta> g) = \<epsilon>\<close>
\]
using the interchange law and two zigzag identities. This shows uniqueness.''
\end{quotation}
\<close>
have 1: "g \<cong> g'"
proof -
have "g \<cong> g \<star> ?b"
using assms hcomp_arr_obj isomorphic_reflexive by auto
also have "... \<cong> g \<star> f \<star> g'"
using assms f_in_hhom g_in_hhom g'_in_hhom E'.counit_in_vhom E'.counit_is_iso
isomorphic_def hcomp_ide_isomorphic isomorphic_symmetric
by (metis E'.counit_simps(5) in_hhomE trg_trg)
also have "... \<cong> ?a \<star> g'"
using assms f_in_hhom g_in_hhom g'_in_hhom ide_g' E'.unit_in_vhom E'.unit_is_iso
isomorphic_def hcomp_isomorphic_ide isomorphic_symmetric
by (metis hcomp_assoc hcomp_isomorphic_ide in_hhomE src_src)
also have "... \<cong> g'"
using assms
by (simp add: E'.antipar(1) hcomp_obj_arr isomorphic_reflexive)
finally show ?thesis by blast
qed
have "f \<star> g' \<cong> ?b"
using E'.counit_is_iso isomorphicI [of \<epsilon>'] by auto
hence 2: "f \<star> g \<cong> ?b"
using assms 1 ide_f hcomp_ide_isomorphic [of f g g'] isomorphic_transitive
isomorphic_symmetric
by (metis in_hhomE)
obtain \<xi> where \<xi>: "\<guillemotleft>\<xi> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> iso \<xi>"
using 2 by auto
have \<xi>_in_hom: "\<guillemotleft>\<xi> : f \<star> g \<Rightarrow> ?b\<guillemotright>" and iso_\<xi>: "iso \<xi>"
using \<xi> by auto
have \<xi>_in_hhom: "\<guillemotleft>\<xi> : ?b \<rightarrow> ?b\<guillemotright>"
using b \<xi>_in_hom
by (metis \<xi> in_hhom_def iso_is_arr obj_simps(2-3) vconn_implies_hpar(1-4))
text \<open>
At the time of this writing, the definition of \<open>\<epsilon>\<close> given on nLab
\cite{nlab-adjoint-equivalence} had an apparent typo:
the expression \<open>f \<star> g \<star> \<xi>\<^sup>-\<^sup>1\<close> should read \<open>\<xi>\<^sup>-\<^sup>1 \<star> f \<star> g\<close>, as we have used here.
\<close>
let ?\<epsilon> = "\<xi> \<cdot> (f \<star> inv \<eta> \<star> g) \<cdot> (inv \<xi> \<star> f \<star> g)"
have \<epsilon>_in_hom: "\<guillemotleft>?\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>f \<star> inv \<eta> \<star> g : f \<star> g \<star> f \<star> g \<Rightarrow> f \<star> g\<guillemotright>"
proof -
have "\<guillemotleft>f \<star> inv \<eta> \<star> g : f \<star> (g \<star> f) \<star> g \<Rightarrow> f \<star> g\<guillemotright>"
proof -
have "\<guillemotleft>f \<star> inv \<eta> \<star> g : f \<star> (g \<star> f) \<star> g \<Rightarrow> f \<star> ?a \<star> g\<guillemotright>"
using assms \<eta>_in_hom iso_\<eta> by (intro hcomp_in_vhom) auto
thus ?thesis
using assms f_in_hhom hcomp_obj_arr by (metis in_hhomE)
qed
moreover have "f \<star> (g \<star> f) \<star> g = f \<star> g \<star> f \<star> g"
using hcomp_assoc by simp
ultimately show ?thesis by simp
qed
show "\<guillemotleft>inv \<xi> \<star> f \<star> g : f \<star> g \<Rightarrow> f \<star> g \<star> f \<star> g\<guillemotright>"
proof -
have "\<guillemotleft>inv \<xi> \<star> f \<star> g : ?b \<star> f \<star> g \<Rightarrow> (f \<star> g) \<star> f \<star> g\<guillemotright>"
using assms \<xi>_in_hom iso_\<xi> by (intro hcomp_in_vhom, auto)
thus ?thesis
using hcomp_assoc f_in_hhom g_in_hhom b hcomp_obj_arr [of ?b "f \<star> g"]
by fastforce
qed
show "\<guillemotleft>\<xi> : f \<star> g \<Rightarrow> ?b\<guillemotright>"
using \<xi>_in_hom by blast
qed
have "iso ?\<epsilon>"
using f_in_hhom g_in_hhom \<eta>_in_hhom ide_f ide_g \<eta>_in_hom iso_\<eta> \<xi>_in_hhom \<xi>_in_hom iso_\<xi>
iso_inv_iso isos_compose
by (metis \<epsilon>_in_hom arrI hseqE ide_is_iso iso_hcomp seqE)
have 4: "\<guillemotleft>inv \<xi> \<star> f : ?b \<star> f \<Rightarrow> f \<star> g \<star> f\<guillemotright>"
proof -
have "\<guillemotleft>inv \<xi> \<star> f : ?b \<star> f \<Rightarrow> (f \<star> g) \<star> f\<guillemotright>"
using \<xi>_in_hom iso_\<xi> f_in_hhom
by (intro hcomp_in_vhom, auto)
thus ?thesis
using hcomp_assoc by simp
qed
text \<open>
First show \<open>?\<epsilon>\<close> and \<open>\<eta>\<close> satisfy the ``left'' triangle identity.
\<close>
have triangle_left: "(?\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f"
proof -
have "(?\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = (\<xi> \<star> f) \<cdot> (f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (inv \<xi> \<star> f \<star> g \<star> f) \<cdot> (?b \<star> f \<star> \<eta>)"
proof -
have "f \<star> \<eta> = ?b \<star> f \<star> \<eta>"
using b \<eta>_in_hhom hcomp_obj_arr [of ?b "f \<star> \<eta>"] by fastforce
moreover have "\<xi> \<cdot> (f \<star> inv \<eta> \<star> g) \<cdot> (inv \<xi> \<star> f \<star> g) \<star> f =
(\<xi> \<star> f) \<cdot> ((f \<star> inv \<eta> \<star> g) \<star> f) \<cdot> ((inv \<xi> \<star> f \<star> g) \<star> f)"
using ide_f ide_g \<xi>_in_hhom \<xi>_in_hom iso_\<xi> \<eta>_in_hhom \<eta>_in_hom iso_\<eta> whisker_right
by (metis \<epsilon>_in_hom arrI seqE)
moreover have "... = (\<xi> \<star> f) \<cdot> (f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (inv \<xi> \<star> f \<star> g \<star> f)"
using hcomp_assoc by simp
ultimately show ?thesis
using comp_assoc by simp
qed
also have "... = (\<xi> \<star> f) \<cdot> ((f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (f \<star> g \<star> f \<star> \<eta>)) \<cdot> (inv \<xi> \<star> f)"
proof -
have "(inv \<xi> \<star> f \<star> g \<star> f) \<cdot> (?b \<star> f \<star> \<eta>) = ((inv \<xi> \<star> f) \<star> (g \<star> f)) \<cdot> ((?b \<star> f) \<star> \<eta>)"
using hcomp_assoc by simp
also have "... = (inv \<xi> \<star> f) \<cdot> (?b \<star> f) \<star> (g \<star> f) \<cdot> \<eta>"
proof -
have "seq (inv \<xi> \<star> f) (?b \<star> f)"
using a b 4 ide_f ide_g \<xi>_in_hhom \<xi>_in_hom iso_\<xi> by blast
moreover have "seq (g \<star> f) \<eta>"
using f_in_hhom g_in_hhom ide_g ide_f \<eta>_in_hom by fast
ultimately show ?thesis
using interchange [of "inv \<xi> \<star> f" "?b \<star> f" "g \<star> f" \<eta>] by simp
qed
also have "... = inv \<xi> \<star> f \<star> \<eta>"
using 4 comp_arr_dom comp_cod_arr \<eta>_in_hom hcomp_assoc by (metis in_homE)
also have "... = (f \<star> g) \<cdot> inv \<xi> \<star> (f \<star> \<eta>) \<cdot> (f \<star> ?a)"
proof -
have "(f \<star> g) \<cdot> inv \<xi> = inv \<xi>"
using \<xi>_in_hom iso_\<xi> comp_cod_arr by auto
moreover have "(f \<star> \<eta>) \<cdot> (f \<star> ?a) = f \<star> \<eta>"
proof -
have "\<guillemotleft>f \<star> \<eta> : f \<star> ?a \<Rightarrow> f \<star> g \<star> f\<guillemotright>"
using \<eta>_in_hom by fastforce
thus ?thesis
using comp_arr_dom by blast
qed
ultimately show ?thesis by argo
qed
also have "... = ((f \<star> g) \<star> (f \<star> \<eta>)) \<cdot> (inv \<xi> \<star> (f \<star> ?a))"
proof -
have "seq (f \<star> g) (inv \<xi>)"
using \<xi>_in_hom iso_\<xi> comp_cod_arr by auto
moreover have "seq (f \<star> \<eta>) (f \<star> ?a)"
using f_in_hhom \<eta>_in_hom by force
ultimately show ?thesis
using interchange by simp
qed
also have "... = (f \<star> g \<star> f \<star> \<eta>) \<cdot> (inv \<xi> \<star> f)"
using hcomp_arr_obj hcomp_assoc by auto
finally have "(inv \<xi> \<star> f \<star> g \<star> f) \<cdot> (?b \<star> f \<star> \<eta>) = (f \<star> g \<star> f \<star> \<eta>) \<cdot> (inv \<xi> \<star> f)"
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (\<xi> \<star> f) \<cdot> ((f \<star> ?a \<star> \<eta>) \<cdot> (f \<star> inv \<eta> \<star> ?a)) \<cdot> (inv \<xi> \<star> f)"
proof -
have "(f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (f \<star> (g \<star> f) \<star> \<eta>) = f \<star> (inv \<eta> \<star> g \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>)"
proof -
have "(f \<star> (inv \<eta> \<star> g) \<star> f) \<cdot> (f \<star> (g \<star> f) \<star> \<eta>) = f \<star> ((inv \<eta> \<star> g) \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>)"
proof -
have "seq ((inv \<eta> \<star> g) \<star> f) ((g \<star> f) \<star> \<eta>)"
proof -
have "seq (inv \<eta> \<star> g \<star> f) ((g \<star> f) \<star> \<eta>)"
using f_in_hhom ide_f g_in_hhom ide_g \<eta>_in_hhom \<eta>_in_hom iso_\<eta>
apply (intro seqI hseqI')
apply auto
by fastforce+
thus ?thesis
using hcomp_assoc by simp
qed
thus ?thesis
using whisker_left by simp
qed
thus ?thesis
using hcomp_assoc by simp
qed
also have "... = f \<star> (?a \<star> \<eta>) \<cdot> (inv \<eta> \<star> ?a)"
proof -
have "(inv \<eta> \<star> g \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>) = (?a \<star> \<eta>) \<cdot> (inv \<eta> \<star> ?a)"
proof -
have "(inv \<eta> \<star> g \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>) = inv \<eta> \<cdot> (g \<star> f) \<star> (g \<star> f) \<cdot> \<eta>"
using g_in_hhom ide_g \<eta>_in_hom iso_\<eta>
interchange [of "inv \<eta>" "g \<star> f" "g \<star> f" \<eta>]
by force
also have "... = inv \<eta> \<star> \<eta>"
using \<eta>_in_hom iso_\<eta> comp_arr_dom comp_cod_arr by auto
also have "... = ?a \<cdot> inv \<eta> \<star> \<eta> \<cdot> ?a"
using \<eta>_in_hom iso_\<eta> comp_arr_dom comp_cod_arr by auto
also have "... = (?a \<star> \<eta>) \<cdot> (inv \<eta> \<star> ?a)"
using a \<eta>_in_hom iso_\<eta> interchange [of ?a "inv \<eta>" \<eta> ?a] by blast
finally show ?thesis by simp
qed
thus ?thesis by argo
qed
also have "... = (f \<star> ?a \<star> \<eta>) \<cdot> (f \<star> inv \<eta> \<star> ?a)"
proof -
have "seq (?a \<star> \<eta>) (inv \<eta> \<star> ?a)"
proof (intro seqI')
show "\<guillemotleft>inv \<eta> \<star> ?a : (g \<star> f) \<star> ?a \<Rightarrow> ?a \<star> ?a\<guillemotright>"
using a g_in_hhom \<eta>_in_hom iso_\<eta> hseqI ide_f ide_g
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>?a \<star> \<eta> : ?a \<star> ?a \<Rightarrow> ?a \<star> g \<star> f\<guillemotright>"
using a \<eta>_in_hom hseqI by (intro hcomp_in_vhom) auto
qed
thus ?thesis
using whisker_left by simp
qed
finally show ?thesis
using hcomp_assoc by simp
qed
also have "... = (\<xi> \<star> f) \<cdot> ((f \<star> \<eta>) \<cdot> (f \<star> inv \<eta>)) \<cdot> (inv \<xi> \<star> f)"
using a \<eta>_in_hhom iso_\<eta> hcomp_obj_arr [of ?a \<eta>] hcomp_arr_obj [of "inv \<eta>" ?a] by auto
also have "... = (\<xi> \<star> f) \<cdot> (inv \<xi> \<star> f)"
proof -
have "((f \<star> \<eta>) \<cdot> (f \<star> inv \<eta>)) \<cdot> (inv \<xi> \<star> f) = (f \<star> \<eta> \<cdot> inv \<eta>) \<cdot> (inv \<xi> \<star> f)"
using \<eta>_in_hhom iso_\<eta> whisker_left inv_in_hom by auto
also have "... = (f \<star> g \<star> f) \<cdot> (inv \<xi> \<star> f)"
using \<eta>_in_hom iso_\<eta> comp_arr_inv inv_is_inverse by auto
also have "... = inv \<xi> \<star> f"
using 4 comp_cod_arr by blast
ultimately show ?thesis by simp
qed
also have "... = f"
proof -
have "(\<xi> \<star> f) \<cdot> (inv \<xi> \<star> f) = \<xi> \<cdot> inv \<xi> \<star> f"
using \<xi>_in_hhom iso_\<xi> whisker_right by auto
also have "... = ?b \<star> f"
using \<xi>_in_hom iso_\<xi> comp_arr_inv' by auto
also have "... = f"
using hcomp_obj_arr by auto
finally show ?thesis by blast
qed
finally show ?thesis by blast
qed
(* TODO: Putting this earlier breaks some steps in the proof. *)
interpret E: equivalence_in_strict_bicategory V H \<a> \<i> src trg f g \<eta> ?\<epsilon>
using ide_g \<eta>_in_hom \<epsilon>_in_hom g_in_hhom \<open>iso \<eta>\<close> \<open>iso ?\<epsilon>\<close>
by (unfold_locales, auto)
text \<open>
Apply ``triangle left if and only iff right'' to show the ``right'' triangle identity.
\<close>
have triangle_right: "((g \<star> \<xi> \<cdot> (f \<star> inv \<eta> \<star> g) \<cdot> (inv \<xi> \<star> f \<star> g)) \<cdot> (\<eta> \<star> g) = g)"
using triangle_left E.triangle_left_iff_right by simp
text \<open>
Use the two triangle identities to establish an adjoint equivalence and show that
there is only one choice for the counit.
\<close>
show "\<exists>!\<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
have "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> ?\<epsilon>"
proof
show "(?\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "(?\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = (?\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>)"
proof -
have "seq \<a>\<^sup>-\<^sup>1[f, g, f] (f \<star> \<eta>)"
using E.antipar
by (intro seqI, auto)
hence "\<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = f \<star> \<eta>"
using ide_f ide_g E.antipar triangle_right strict_assoc' comp_ide_arr
by presburger
thus ?thesis by simp
qed
also have "... = f"
using triangle_left by simp
also have "... = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
using strict_lunit strict_runit by simp
finally show ?thesis by simp
qed
show "(g \<star> ?\<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
proof -
have "(g \<star> ?\<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = (g \<star> ?\<epsilon>) \<cdot> (\<eta> \<star> g)"
proof -
have "seq \<a>[g, f, g] (\<eta> \<star> g)"
using E.antipar
by (intro seqI, auto)
hence "\<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<eta> \<star> g"
using ide_f ide_g E.antipar triangle_right strict_assoc comp_ide_arr
by presburger
thus ?thesis by simp
qed
also have "... = g"
using triangle_right by simp
also have "... = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using strict_lunit strict_runit by simp
finally show ?thesis by blast
qed
qed
moreover have "\<And>\<epsilon> \<epsilon>'. \<lbrakk> adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>;
adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>' \<rbrakk>
\<Longrightarrow> \<epsilon> = \<epsilon>'"
using adjunction_unit_determines_counit
by (meson adjoint_equivalence_in_bicategory.axioms(2))
ultimately show ?thesis by auto
qed
qed
end
text \<open>
We now apply strictification to generalize the preceding result to an arbitrary bicategory.
\<close>
context bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: fully_faithful_functor V S.vcomp S.UP
using S.UP_is_fully_faithful_functor by auto
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> 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 \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
..
lemma equivalence_refines_to_adjoint_equivalence:
assumes "equivalence_map f" and "\<guillemotleft>g : trg f \<rightarrow> src f\<guillemotright>" and "ide g"
and "\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>" and "iso \<eta>"
shows "\<exists>!\<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
text \<open>
To unpack the consequences of the assumptions, we need to obtain an
interpretation of @{locale equivalence_in_bicategory}, even though we don't
need the associated data other than \<open>f\<close>, \<open>a\<close>, and \<open>b\<close>.
\<close>
obtain g' \<phi> \<psi> where E: "equivalence_in_bicategory V H \<a> \<i> src trg f g' \<phi> \<psi>"
using assms equivalence_map_def by auto
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g' \<phi> \<psi>
using E by auto
let ?a = "src f" and ?b = "trg f"
have ide_f: "ide f" by simp
have f_in_hhom: "\<guillemotleft>f : ?a \<rightarrow> ?b\<guillemotright>" by simp
have a: "obj ?a" and b: "obj ?b" by auto
have 1: "S.equivalence_map (S.UP f)"
using assms UP.preserves_equivalence_maps by simp
let ?\<eta>' = "S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a"
have 2: "\<guillemotleft>S.UP \<eta> : S.UP ?a \<Rightarrow>\<^sub>S S.UP (g \<star> f)\<guillemotright>"
using assms UP.preserves_hom [of \<eta> "src f" "g \<star> f"] by auto
have 3: "\<guillemotleft>?\<eta>' : UP.map\<^sub>0 ?a \<Rightarrow>\<^sub>S S.UP g \<star>\<^sub>S S.UP f\<guillemotright> \<and> S.iso ?\<eta>'"
proof (intro S.comp_in_homI conjI)
show "\<guillemotleft>S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) : S.UP (g \<star> f) \<Rightarrow>\<^sub>S S.UP g \<star>\<^sub>S S.UP f\<guillemotright>"
using assms UP.cmp_in_hom(2) by auto
moreover show "\<guillemotleft>UP.unit ?a : UP.map\<^sub>0 ?a \<Rightarrow>\<^sub>S S.UP ?a\<guillemotright>" by auto
moreover show "\<guillemotleft>S.UP \<eta> : S.UP ?a \<Rightarrow>\<^sub>S S.UP (g \<star> f)\<guillemotright>"
using 2 by simp
ultimately show "S.iso (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a)"
using assms UP.cmp_components_are_iso UP.unit_char(2)
by (intro S.isos_compose) auto
qed
have ex_un_\<xi>': "\<exists>!\<xi>'. adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
(S.UP f) (S.UP g) ?\<eta>' \<xi>'"
proof -
have "\<guillemotleft>S.UP g : S.trg (S.UP f) \<rightarrow>\<^sub>S S.src (S.UP f)\<guillemotright>"
using assms(2) by auto
moreover have "S.ide (S.UP g)"
by (simp add: assms(3))
ultimately show ?thesis
using 1 3 S.equivalence_refines_to_adjoint_equivalence S.UP_map\<^sub>0_obj by simp
qed
obtain \<xi>' where \<xi>': "adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
(S.UP f) (S.UP g) ?\<eta>' \<xi>'"
using ex_un_\<xi>' by auto
interpret E': adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
\<open>S.UP f\<close> \<open>S.UP g\<close> ?\<eta>' \<xi>'
using \<xi>' by auto
let ?\<epsilon>' = "UP.unit ?b \<cdot>\<^sub>S \<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g))"
have \<epsilon>': "\<guillemotleft>?\<epsilon>' : S.UP (f \<star> g) \<Rightarrow>\<^sub>S S.UP ?b\<guillemotright>"
using assms(2-3) by auto
have ex_un_\<epsilon>: "\<exists>!\<epsilon>. \<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> S.UP \<epsilon> = ?\<epsilon>'"
proof -
have "\<exists>\<epsilon>. \<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> S.UP \<epsilon> = ?\<epsilon>'"
proof -
have "src (f \<star> g) = src ?b \<and> trg (f \<star> g) = trg ?b"
using assms(2) f_in_hhom by auto
moreover have "ide (f \<star> g)"
using assms(2-3) by auto
ultimately show ?thesis
using \<epsilon>' UP.locally_full by auto
qed
moreover have
"\<And>\<mu> \<nu>. \<lbrakk> \<guillemotleft>\<mu> : f \<star> g \<Rightarrow> ?b\<guillemotright>; S.UP \<mu> = ?\<epsilon>'; \<guillemotleft>\<nu> : f \<star> g \<Rightarrow> ?b\<guillemotright>; S.UP \<nu> = ?\<epsilon>' \<rbrakk>
\<Longrightarrow> \<mu> = \<nu>"
proof -
fix \<mu> \<nu>
assume \<mu>: "\<guillemotleft>\<mu> : f \<star> g \<Rightarrow> ?b\<guillemotright>" and \<nu>: "\<guillemotleft>\<nu> : f \<star> g \<Rightarrow> ?b\<guillemotright>"
and 1: "S.UP \<mu> = ?\<epsilon>'" and 2: "S.UP \<nu> = ?\<epsilon>'"
have "par \<mu> \<nu>"
using \<mu> \<nu> by fastforce
thus "\<mu> = \<nu>"
using 1 2 UP.is_faithful [of \<mu> \<nu>] by simp
qed
ultimately show ?thesis by auto
qed
have iso_\<epsilon>': "S.iso ?\<epsilon>'"
proof (intro S.isos_compose)
show "S.iso (S.inv (S.cmp\<^sub>U\<^sub>P (f, g)))"
using assms UP.cmp_components_are_iso by auto
show "S.iso \<xi>'"
using E'.counit_is_iso by blast
show "S.iso (UP.unit ?b)"
using b UP.unit_char(2) by simp
show "S.seq (UP.unit ?b) (\<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)))"
proof (intro S.seqI')
show "\<guillemotleft>UP.unit ?b : UP.map\<^sub>0 ?b \<Rightarrow>\<^sub>S S.UP ?b\<guillemotright>"
using b UP.unit_char by simp
show "\<guillemotleft>\<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) : S.UP (f \<star> g) \<Rightarrow>\<^sub>S UP.map\<^sub>0 ?b\<guillemotright>"
using assms by auto
qed
thus "S.seq \<xi>' (S.inv (S.cmp\<^sub>U\<^sub>P (f, g)))" by auto
qed
obtain \<epsilon> where \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> S.UP \<epsilon> = ?\<epsilon>'"
using ex_un_\<epsilon> by auto
interpret E'': equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms(1,3-5)
apply unfold_locales
apply simp_all
using assms(2) \<epsilon>
apply auto[1]
using \<epsilon> iso_\<epsilon>' UP.reflects_iso [of \<epsilon>]
by auto
interpret E'': adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
proof
show "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) =
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
using E''.UP_triangle(3) by simp
also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S
(UP.unit ?b \<cdot>\<^sub>S \<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
using \<epsilon> S.comp_assoc by simp
also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<cdot>\<^sub>S \<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "\<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) = \<xi>'"
proof -
have "S.iso (S.cmp\<^sub>U\<^sub>P (f, g))"
using assms by auto
moreover have "S.dom (S.cmp\<^sub>U\<^sub>P (f, g)) = S.UP f \<star>\<^sub>S S.UP g"
using assms by auto
ultimately have "S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) = S.UP f \<star>\<^sub>S S.UP g"
using S.comp_inv_arr' by simp
thus ?thesis
using S.comp_arr_dom E'.counit_in_hom(2) by simp
qed
thus ?thesis by argo
qed
also have
"... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "UP.unit ?b \<cdot>\<^sub>S \<xi>' \<star>\<^sub>S S.UP f = (UP.unit ?b \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (\<xi>' \<star>\<^sub>S S.UP f)"
using assms b UP.unit_char S.whisker_right S.UP_map\<^sub>0_obj by auto
moreover have "S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> =
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)"
using assms S.whisker_left by auto
ultimately show ?thesis
using S.comp_assoc by simp
qed
also have "... = (S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)
= (S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have "(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>) =
S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>"
using assms S.whisker_left by auto
hence "((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>))
= ((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>))"
by simp
also have "... = ((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) = \<xi>' \<star>\<^sub>S S.UP f"
proof -
have "\<guillemotleft>\<xi>' \<star>\<^sub>S S.UP f :
(S.UP f \<star>\<^sub>S S.UP g) \<star>\<^sub>S S.UP f \<Rightarrow>\<^sub>S S.trg (S.UP f) \<star>\<^sub>S S.UP f\<guillemotright>"
using assms by auto
moreover have "\<guillemotleft>S.\<a>' (S.UP f) (S.UP g) (S.UP f) :
S.UP f \<star>\<^sub>S S.UP g \<star>\<^sub>S S.UP f \<Rightarrow>\<^sub>S (S.UP f \<star>\<^sub>S S.UP g) \<star>\<^sub>S S.UP f\<guillemotright>"
using assms S.assoc'_in_hom by auto
ultimately show ?thesis
using assms S.strict_assoc' S.iso_assoc S.hcomp_assoc E'.antipar
S.comp_arr_ide S.seqI'
by (metis (no_types, lifting) E'.ide_left E'.ide_right)
qed
thus ?thesis
using S.comp_assoc by simp
qed
also have "... = ((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>))"
using S.comp_assoc by simp
also have "... = (S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) =
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) =
S.lunit' (S.UP f) \<cdot>\<^sub>S S.runit (S.UP f)"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) =
(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a)"
proof -
have "S.seq (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) (UP.unit ?a)"
using assms UP.unit_char UP.cmp_components_are_iso
E'.unit_simps(1) S.comp_assoc
by presburger
hence "(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) =
S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a"
using assms UP.unit_char UP.cmp_components_are_iso S.comp_assoc
S.whisker_left [of "S.UP f" "S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>" "UP.unit ?a"]
by simp
thus ?thesis by simp
qed
thus ?thesis
using assms E'.triangle_left UP.cmp_components_are_iso UP.unit_char
by argo
qed
also have "... = S.UP f"
using S.strict_lunit' S.strict_runit by simp
finally have 1: "((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) = S.UP f"
using S.comp_assoc by simp
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) =
S.UP f \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (UP.unit ?a))"
proof -
have "S.arr (S.UP f)"
using assms by simp
moreover have "S.iso (S.UP f \<star>\<^sub>S UP.unit ?a)"
using assms UP.unit_char S.UP_map\<^sub>0_obj by auto
moreover have "S.inv (S.UP f \<star>\<^sub>S UP.unit ?a) =
S.UP f \<star>\<^sub>S S.inv (UP.unit ?a)"
using assms a UP.unit_char S.UP_map\<^sub>0_obj by auto
ultimately show ?thesis
using assms 1 UP.unit_char UP.cmp_components_are_iso
S.invert_side_of_triangle(2)
[of "S.UP f" "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)"
"S.UP f \<star>\<^sub>S UP.unit ?a"]
by presburger
qed
also have "... = S.UP f \<star>\<^sub>S S.inv (UP.unit ?a)"
proof -
have "\<guillemotleft>S.UP f \<star>\<^sub>S S.inv (UP.unit ?a) :
S.UP f \<star>\<^sub>S S.UP ?a \<Rightarrow>\<^sub>S S.UP f \<star>\<^sub>S UP.map\<^sub>0 ?a\<guillemotright>"
using assms ide_f f_in_hhom UP.unit_char [of ?a] S.inv_in_hom
apply (intro S.hcomp_in_vhom)
apply auto[1]
apply blast
by auto
moreover have "S.UP f \<star>\<^sub>S UP.map\<^sub>0 ?a = S.UP f"
using a S.hcomp_arr_obj S.UP_map\<^sub>0_obj by auto
finally show ?thesis
using S.comp_cod_arr by blast
qed
finally show ?thesis by auto
qed
thus ?thesis
using S.comp_assoc by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using S.comp_assoc by simp
qed
also have "... = S.UP \<l>\<^sup>-\<^sup>1[f] \<cdot>\<^sub>S S.UP \<r>[f]"
proof -
have "S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<star>\<^sub>S S.UP f) = S.UP \<l>\<^sup>-\<^sup>1[f]"
proof -
have "S.UP f = S.UP \<l>[f] \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)"
using UP.lunit_coherence iso_lunit S.strict_lunit by simp
thus ?thesis
using UP.image_of_unitor(3) ide_f by presburger
qed
moreover have "(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))
= S.UP \<r>[f]"
proof -
have "S.UP \<r>[f] \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f)) = S.UP f"
using UP.runit_coherence [of f] S.strict_runit by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f)))"
- using UP.unit_char UP.cmp_components_are_iso VV.arr_char S.UP_map\<^sub>0_obj
+ using UP.unit_char UP.cmp_components_are_iso VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.UP_map\<^sub>0_obj
by (intro S.isos_compose) auto
ultimately have
"S.UP \<r>[f] = S.UP f \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f)))"
using S.invert_side_of_triangle(2)
[of "S.UP f" "S.UP \<r>[f]" "S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f))"]
ideD(1) ide_f by blast
thus ?thesis
using ide_f UP.image_of_unitor(2) [of f] by argo
qed
ultimately show ?thesis
using S.comp_assoc by simp
qed
also have "... = S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
by simp
finally have "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) = S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
by simp
moreover have "par ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
proof -
have "\<guillemotleft>(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) : f \<star> src f \<Rightarrow> trg f \<star> f\<guillemotright>"
using E''.triangle_in_hom(1) by simp
moreover have "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] : f \<star> src f \<Rightarrow> trg f \<star> f\<guillemotright>" by auto
ultimately show ?thesis
by (metis in_homE)
qed
ultimately show ?thesis
using UP.is_faithful by blast
qed
thus "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using E''.triangle_left_implies_right by simp
qed
show ?thesis
using E''.adjoint_equivalence_in_bicategory_axioms E''.adjunction_in_bicategory_axioms
adjunction_unit_determines_counit adjoint_equivalence_in_bicategory_def
by metis
qed
lemma equivalence_map_extends_to_adjoint_equivalence:
assumes "equivalence_map f"
shows "\<exists>g \<eta> \<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
obtain g \<eta> \<epsilon>' where E: "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'"
using assms equivalence_map_def by auto
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using E by auto
obtain \<epsilon> where A: "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms equivalence_refines_to_adjoint_equivalence [of f g \<eta>]
E.antipar E.unit_is_iso E.unit_in_hom by auto
show ?thesis
using E A by blast
qed
end
subsection "Uniqueness of Adjoints"
text \<open>
Left and right adjoints determine each other up to isomorphism.
\<close>
context strict_bicategory
begin
lemma left_adjoint_determines_right_up_to_iso:
assumes "adjoint_pair f g" and "adjoint_pair f g'"
shows "g \<cong> g'"
proof -
obtain \<eta> \<epsilon> where A: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using A by auto
interpret A: adjunction_in_strict_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon> ..
obtain \<eta>' \<epsilon>' where A': "adjunction_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret A': adjunction_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'
using A' by auto
interpret A': adjunction_in_strict_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>' ..
let ?\<phi> = "A'.trnl\<^sub>\<eta> g \<epsilon>"
have "\<guillemotleft>?\<phi>: g \<Rightarrow> g'\<guillemotright>"
using A'.trnl\<^sub>\<eta>_eq A'.adjoint_transpose_left(1) [of "trg f" g] A.antipar A'.antipar
hcomp_arr_obj
by auto
moreover have "iso ?\<phi>"
proof (intro isoI)
let ?\<psi> = "A.trnl\<^sub>\<eta> g' \<epsilon>'"
show "inverse_arrows ?\<phi> ?\<psi>"
proof
show "ide (?\<phi> \<cdot> ?\<psi>)"
proof -
have 1: "ide (trg f) \<and> trg (trg f) = trg f"
by simp
have "?\<phi> \<cdot> ?\<psi> = (g' \<star> \<epsilon>) \<cdot> ((\<eta>' \<star> g) \<cdot> (g \<star> \<epsilon>')) \<cdot> (\<eta> \<star> g')"
using 1 A.antipar A'.antipar A.trnl\<^sub>\<eta>_eq [of "trg f" g' \<epsilon>']
A'.trnl\<^sub>\<eta>_eq [of "trg f" g \<epsilon>] comp_assoc A.counit_in_hom A'.counit_in_hom
by simp
also have "... = ((g' \<star> \<epsilon>) \<cdot> (g' \<star> f \<star> g \<star> \<epsilon>')) \<cdot> ((\<eta>' \<star> g \<star> f \<star> g') \<cdot> (\<eta> \<star> g'))"
proof -
have "(\<eta>' \<star> g) \<cdot> (g \<star> \<epsilon>') = (\<eta>' \<star> g \<star> trg f) \<cdot> (src f \<star> g \<star> \<epsilon>')"
using A.antipar A'.antipar hcomp_arr_obj hcomp_obj_arr [of "src f" "g \<star> \<epsilon>'"]
hseqI'
by (metis A'.counit_simps(1) A'.counit_simps(5) A.ide_right ideD(1)
obj_trg trg_hcomp)
also have "... = \<eta>' \<star> g \<star> \<epsilon>'"
using A.antipar A'.antipar interchange [of \<eta>' "src f" "g \<star> trg f" "g \<star> \<epsilon>'"]
whisker_left comp_arr_dom comp_cod_arr
by simp
also have "... = ((g' \<star> f) \<star> g \<star> \<epsilon>') \<cdot> (\<eta>' \<star> g \<star> (f \<star> g'))"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
A'.unit_in_hom A'.counit_in_hom interchange whisker_left
comp_arr_dom comp_cod_arr
by (metis A'.counit_simps(1-2,5) A'.unit_simps(1,3) hseqI' ide_char)
also have "... = (g' \<star> f \<star> g \<star> \<epsilon>') \<cdot> (\<eta>' \<star> g \<star> f \<star> g')"
using hcomp_assoc by simp
finally show ?thesis
using comp_assoc by simp
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> ((g' \<star> (\<epsilon> \<star> f) \<star> g') \<cdot> (g' \<star> (f \<star> \<eta>) \<star> g')) \<cdot> (\<eta>' \<star> g')"
proof -
have "(g' \<star> \<epsilon>) \<cdot> (g' \<star> f \<star> g \<star> \<epsilon>') = (g' \<star> \<epsilon>') \<cdot> (g' \<star> \<epsilon> \<star> f \<star> g')"
proof -
have "(g' \<star> \<epsilon>) \<cdot> (g' \<star> f \<star> g \<star> \<epsilon>') = g' \<star> \<epsilon> \<star> \<epsilon>'"
proof -
have "\<epsilon> \<cdot> (f \<star> g \<star> \<epsilon>') = \<epsilon> \<star> \<epsilon>'"
using A.ide_left A.ide_right A.antipar A'.antipar hcomp_arr_obj comp_arr_dom
comp_cod_arr interchange obj_src trg_src
by (metis A'.counit_simps(1,3) A.counit_simps(1-2,4) hcomp_assoc)
thus ?thesis
using A.antipar A'.antipar whisker_left [of g' \<epsilon> "f \<star> g \<star> \<epsilon>'"]
by (simp add: hcomp_assoc)
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> (g' \<star> \<epsilon> \<star> f \<star> g')"
proof -
have "\<epsilon> \<star> \<epsilon>' = \<epsilon>' \<cdot> (\<epsilon> \<star> f \<star> g')"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
hcomp_obj_arr hcomp_arr_obj comp_arr_dom comp_cod_arr interchange
obj_src trg_src
by (metis A'.counit_simps(1-2,5) A.counit_simps(1,3-4) arr_cod
not_arr_null seq_if_composable)
thus ?thesis
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
whisker_left
by (metis A'.counit_simps(1,5) A.counit_simps(1,4) hseqI')
qed
finally show ?thesis by simp
qed
moreover have "(\<eta>' \<star> g \<star> f \<star> g') \<cdot> (\<eta> \<star> g') = (g' \<star> f \<star> \<eta> \<star> g') \<cdot> (\<eta>' \<star> g')"
proof -
have "(\<eta>' \<star> g \<star> f \<star> g') \<cdot> (\<eta> \<star> g') = \<eta>' \<star> \<eta> \<star> g'"
proof -
have "(\<eta>' \<star> g \<star> f) \<cdot> \<eta> = \<eta>' \<star> \<eta>"
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj
interchange comp_arr_dom comp_cod_arr
by (metis A'.unit_simps(1-2,4) A.unit_simps(1,3,5) hcomp_obj_arr obj_trg)
thus ?thesis
using A.antipar A'.antipar whisker_right [of g' "\<eta>' \<star> g \<star> f" \<eta>]
by (simp add: hcomp_assoc)
qed
also have "... = (g' \<star> f \<star> \<eta> \<star> g') \<cdot> (\<eta>' \<star> g')"
proof -
have "\<eta>' \<star> \<eta> = (g' \<star> f \<star> \<eta>) \<cdot> \<eta>'"
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj
comp_arr_dom comp_cod_arr hcomp_assoc interchange
by (metis A'.unit_simps(1,3-4) A.unit_simps(1-2) obj_src)
thus ?thesis
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj
whisker_right [of g' "g' \<star> f \<star> \<eta>" \<eta>']
by (metis A'.ide_right A'.unit_simps(1,4) A.unit_simps(1,5)
hseqI' hcomp_assoc)
qed
finally show ?thesis by simp
qed
ultimately show ?thesis
using comp_assoc hcomp_assoc by simp
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> ((g' \<star> f) \<star> g') \<cdot> (\<eta>' \<star> g')"
proof -
have "(g' \<star> (\<epsilon> \<star> f) \<star> g') \<cdot> (g' \<star> (f \<star> \<eta>) \<star> g') = g' \<star> f \<star> g'"
proof -
have "(g' \<star> (\<epsilon> \<star> f) \<star> g') \<cdot> (g' \<star> (f \<star> \<eta>) \<star> g') =
g' \<star> ((\<epsilon> \<star> f) \<star> g') \<cdot> ((f \<star> \<eta>) \<star> g')"
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom
A'.counit_in_hom whisker_left [of g' "(\<epsilon> \<star> f) \<star> g'" "(f \<star> \<eta>) \<star> g'"]
by (metis A'.ide_right A.triangle_left hseqI' ideD(1) whisker_right)
also have "... = g' \<star> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) \<star> g'"
using A.antipar A'.antipar whisker_right [of g' "\<epsilon> \<star> f" "f \<star> \<eta>"]
by (simp add: A.triangle_left)
also have "... = g' \<star> f \<star> g'"
using A.triangle_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using hcomp_assoc by simp
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> (\<eta>' \<star> g')"
using A.antipar A'.antipar A'.unit_in_hom A'.counit_in_hom comp_cod_arr
by (metis A'.ide_right A'.triangle_in_hom(2) A.ide_left arrI assoc_is_natural_2
ide_char seqE strict_assoc)
also have "... = g'"
using A'.triangle_right by simp
finally have "?\<phi> \<cdot> ?\<psi> = g'" by simp
thus ?thesis by simp
qed
show "ide (?\<psi> \<cdot> ?\<phi>)"
proof -
have 1: "ide (trg f) \<and> trg (trg f) = trg f"
by simp
have "?\<psi> \<cdot> ?\<phi> = (g \<star> \<epsilon>') \<cdot> ((\<eta> \<star> g') \<cdot> (g' \<star> \<epsilon>)) \<cdot> (\<eta>' \<star> g)"
using A.antipar A'.antipar A'.trnl\<^sub>\<eta>_eq [of "trg f" g \<epsilon>]
A.trnl\<^sub>\<eta>_eq [of "trg f" g' \<epsilon>'] comp_assoc A.counit_in_hom A'.counit_in_hom
by simp
also have "... = ((g \<star> \<epsilon>') \<cdot> (g \<star> f \<star> g' \<star> \<epsilon>)) \<cdot> ((\<eta> \<star> g' \<star> f \<star> g) \<cdot> (\<eta>' \<star> g))"
proof -
have "(\<eta> \<star> g') \<cdot> (g' \<star> \<epsilon>) = (\<eta> \<star> g' \<star> trg f) \<cdot> (src f \<star> g' \<star> \<epsilon>)"
using A.antipar A'.antipar hcomp_arr_obj hcomp_obj_arr hseqI'
by (metis A'.ide_right A.unit_simps(1,4) hcomp_assoc hcomp_obj_arr
ideD(1) obj_src)
also have "... = \<eta> \<star> g' \<star> \<epsilon>"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom
A.counit_in_hom interchange
by (metis "1" A.counit_simps(5) A.unit_simps(4) hseqI' ide_def ide_in_hom(2)
not_arr_null seqI' src.preserves_ide)
also have "... = ((g \<star> f) \<star> g' \<star> \<epsilon>) \<cdot> (\<eta> \<star> g' \<star> (f \<star> g))"
using A'.ide_right A'.antipar interchange ide_char comp_arr_dom comp_cod_arr hseqI'
by (metis A.counit_simps(1-2,5) A.unit_simps(1,3))
also have "... = (g \<star> f \<star> g' \<star> \<epsilon>) \<cdot> (\<eta> \<star> g' \<star> f \<star> g)"
using hcomp_assoc by simp
finally show ?thesis
using comp_assoc by simp
qed
also have "... = (g \<star> \<epsilon>) \<cdot> ((g \<star> (\<epsilon>' \<star> f) \<star> g) \<cdot> (g \<star> (f \<star> \<eta>') \<star> g)) \<cdot> (\<eta> \<star> g)"
proof -
have "(g \<star> \<epsilon>') \<cdot> (g \<star> f \<star> g' \<star> \<epsilon>) = (g \<star> \<epsilon>) \<cdot> (g \<star> \<epsilon>' \<star> f \<star> g)"
proof -
have "(g \<star> \<epsilon>') \<cdot> (g \<star> f \<star> g' \<star> \<epsilon>) = g \<star> \<epsilon>' \<star> \<epsilon>"
proof -
have "\<epsilon>' \<cdot> (f \<star> g' \<star> \<epsilon>) = \<epsilon>' \<star> \<epsilon>"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_arr_obj
comp_arr_dom comp_cod_arr interchange obj_src trg_src hcomp_assoc
by (metis A.counit_simps(1,3) A'.counit_simps(1-2,4))
thus ?thesis
using A.antipar A'.antipar whisker_left [of g \<epsilon>' "f \<star> g' \<star> \<epsilon>"]
by (simp add: hcomp_assoc)
qed
also have "... = (g \<star> \<epsilon>) \<cdot> (g \<star> \<epsilon>' \<star> f \<star> g)"
proof -
have "\<epsilon>' \<star> \<epsilon> = \<epsilon> \<cdot> (\<epsilon>' \<star> f \<star> g)"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_obj_arr
hcomp_arr_obj comp_arr_dom comp_cod_arr interchange obj_src trg_src
by (metis A.counit_simps(1-2,5) A'.counit_simps(1,3-4)
arr_cod not_arr_null seq_if_composable)
thus ?thesis
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
whisker_left
by (metis A.counit_simps(1,5) A'.counit_simps(1,4) hseqI')
qed
finally show ?thesis by simp
qed
moreover have "(\<eta> \<star> g' \<star> f \<star> g) \<cdot> (\<eta>' \<star> g) = (g \<star> f \<star> \<eta>' \<star> g) \<cdot> (\<eta> \<star> g)"
proof -
have "(\<eta> \<star> g' \<star> f \<star> g) \<cdot> (\<eta>' \<star> g) = \<eta> \<star> \<eta>' \<star> g"
proof -
have "(\<eta> \<star> g' \<star> f) \<cdot> \<eta>' = \<eta> \<star> \<eta>'"
using A.antipar A'.antipar A.unit_in_hom hcomp_arr_obj
comp_arr_dom comp_cod_arr hcomp_obj_arr interchange
by (metis A'.unit_simps(1,3,5) A.unit_simps(1-2,4) obj_trg)
thus ?thesis
using A.antipar A'.antipar whisker_right [of g "\<eta> \<star> g' \<star> f" \<eta>']
by (simp add: hcomp_assoc)
qed
also have "... = ((g \<star> f) \<star> \<eta>' \<star> g) \<cdot> (\<eta> \<star> src f \<star> g)"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom
A'.unit_in_hom comp_arr_dom comp_cod_arr interchange
by (metis A'.unit_simps(1-2,4) A.unit_simps(1,3) hseqI' ide_char)
also have "... = (g \<star> f \<star> \<eta>' \<star> g) \<cdot> (\<eta> \<star> g)"
using A.antipar A'.antipar hcomp_assoc
by (simp add: hcomp_obj_arr)
finally show ?thesis by simp
qed
ultimately show ?thesis
using comp_assoc hcomp_assoc by simp
qed
also have "... = (g \<star> \<epsilon>) \<cdot> ((g \<star> f) \<star> g) \<cdot> (\<eta> \<star> g)"
proof -
have "(g \<star> (\<epsilon>' \<star> f) \<star> g) \<cdot> (g \<star> (f \<star> \<eta>') \<star> g) = g \<star> f \<star> g"
proof -
have "(g \<star> (\<epsilon>' \<star> f) \<star> g) \<cdot> (g \<star> (f \<star> \<eta>') \<star> g) =
g \<star> ((\<epsilon>' \<star> f) \<star> g) \<cdot> ((f \<star> \<eta>') \<star> g)"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom
A.counit_in_hom whisker_left
by (metis A'.triangle_left hseqI' ideD(1) whisker_right)
also have "... = g \<star> (\<epsilon>' \<star> f) \<cdot> (f \<star> \<eta>') \<star> g"
using A.antipar A'.antipar whisker_right [of g "\<epsilon>' \<star> f" "f \<star> \<eta>'"]
by (simp add: A'.triangle_left)
also have "... = g \<star> f \<star> g"
using A'.triangle_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using hcomp_assoc by simp
qed
also have "... = (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g)"
using A.antipar A'.antipar A.unit_in_hom A.counit_in_hom comp_cod_arr
by (metis A.ide_left A.ide_right A.triangle_in_hom(2) arrI assoc_is_natural_2
ide_char seqE strict_assoc)
also have "... = g"
using A.triangle_right by simp
finally have "?\<psi> \<cdot> ?\<phi> = g" by simp
moreover have "ide g"
by simp
ultimately show ?thesis by simp
qed
qed
qed
ultimately show ?thesis
using isomorphic_def by auto
qed
end
text \<open>
We now use strictification to extend to arbitrary bicategories.
\<close>
context bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> 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 \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
..
interpretation UP: fully_faithful_functor V S.vcomp S.UP
using S.UP_is_fully_faithful_functor by auto
lemma left_adjoint_determines_right_up_to_iso:
assumes "adjoint_pair f g" and "adjoint_pair f g'"
shows "g \<cong> g'"
proof -
have 0: "ide g \<and> ide g'"
using assms adjoint_pair_def adjunction_in_bicategory_def
adjunction_data_in_bicategory_def adjunction_data_in_bicategory_axioms_def
by metis
have 1: "S.adjoint_pair (S.UP f) (S.UP g) \<and> S.adjoint_pair (S.UP f) (S.UP g')"
using assms UP.preserves_adjoint_pair by simp
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu> : S.UP g \<Rightarrow>\<^sub>S S.UP g'\<guillemotright> \<and> S.iso \<nu>"
using 1 S.left_adjoint_determines_right_up_to_iso S.isomorphic_def by blast
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : g \<Rightarrow> g'\<guillemotright> \<and> S.UP \<mu> = \<nu>"
using 0 \<nu> UP.is_full [of g' g \<nu>] by auto
have "\<guillemotleft>\<mu> : g \<Rightarrow> g'\<guillemotright> \<and> iso \<mu>"
using \<mu> \<nu> UP.reflects_iso by auto
thus ?thesis
using isomorphic_def by auto
qed
lemma right_adjoint_determines_left_up_to_iso:
assumes "adjoint_pair f g" and "adjoint_pair f' g"
shows "f \<cong> f'"
proof -
obtain \<eta> \<epsilon> where A: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using A by auto
obtain \<eta>' \<epsilon>' where A': "adjunction_in_bicategory V H \<a> \<i> src trg f' g \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret A': adjunction_in_bicategory V H \<a> \<i> src trg f' g \<eta>' \<epsilon>'
using A' by auto
interpret Cop: op_bicategory V H \<a> \<i> src trg ..
interpret Aop: adjunction_in_bicategory V Cop.H Cop.\<a> \<i> Cop.src Cop.trg g f \<eta> \<epsilon>
using A.antipar A.triangle_left A.triangle_right Cop.assoc_ide_simp
Cop.lunit_ide_simp Cop.runit_ide_simp
by (unfold_locales, auto)
interpret Aop': adjunction_in_bicategory V Cop.H Cop.\<a> \<i> Cop.src Cop.trg g f' \<eta>' \<epsilon>'
using A'.antipar A'.triangle_left A'.triangle_right Cop.assoc_ide_simp
Cop.lunit_ide_simp Cop.runit_ide_simp
by (unfold_locales, auto)
show ?thesis
using Aop.adjunction_in_bicategory_axioms Aop'.adjunction_in_bicategory_axioms
Cop.left_adjoint_determines_right_up_to_iso Cop.adjoint_pair_def
by blast
qed
end
context chosen_right_adjoints
begin
lemma isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint:
assumes "is_left_adjoint f" and "f \<cong> h"
shows "f\<^sup>* \<cong> h\<^sup>*"
proof -
have 1: "adjoint_pair f f\<^sup>*"
using assms left_adjoint_extends_to_adjoint_pair by blast
moreover have "adjoint_pair h f\<^sup>*"
using assms 1 adjoint_pair_preserved_by_iso isomorphic_symmetric isomorphic_reflexive
by (meson isomorphic_def right_adjoint_simps(1))
thus ?thesis
using left_adjoint_determines_right_up_to_iso left_adjoint_extends_to_adjoint_pair
by blast
qed
end
context bicategory
begin
lemma equivalence_is_adjoint:
assumes "equivalence_map f"
shows equivalence_is_left_adjoint: "is_left_adjoint f"
and equivalence_is_right_adjoint: "is_right_adjoint f"
proof -
obtain g \<eta> \<epsilon> where fg: "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms equivalence_map_extends_to_adjoint_equivalence by blast
interpret fg: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by simp
interpret gf: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg g f \<open>inv \<epsilon>\<close> \<open>inv \<eta>\<close>
using fg.dual_adjoint_equivalence by simp
show "is_left_adjoint f"
using fg.adjunction_in_bicategory_axioms adjoint_pair_def by auto
show "is_right_adjoint f"
using gf.adjunction_in_bicategory_axioms adjoint_pair_def by auto
qed
lemma right_adjoint_to_equivalence_is_equivalence:
assumes "equivalence_map f" and "adjoint_pair f g"
shows "equivalence_map g"
proof -
obtain \<eta> \<epsilon> where fg: "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret fg: adjunction_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by simp
obtain g' \<phi> \<psi> where fg': "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g' \<phi> \<psi>"
using assms equivalence_map_def by auto
interpret fg': equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg f g' \<phi> \<psi>
using fg' by auto
obtain \<psi>' where \<psi>': "adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g' \<phi> \<psi>'"
using assms equivalence_refines_to_adjoint_equivalence [of f g' \<phi>]
fg'.antipar fg'.unit_in_hom fg'.unit_is_iso
by auto
interpret \<psi>': adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg f g' \<phi> \<psi>'
using \<psi>' by simp
have 1: "g \<cong> g'"
using fg.adjunction_in_bicategory_axioms \<psi>'.adjunction_in_bicategory_axioms
left_adjoint_determines_right_up_to_iso adjoint_pair_def
by blast
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : g' \<Rightarrow> g\<guillemotright> \<and> iso \<gamma>"
using 1 isomorphic_def isomorphic_symmetric by metis
have "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g ((\<gamma> \<star> f) \<cdot> \<phi>) (\<psi>' \<cdot> (f \<star> inv \<gamma>))"
using \<gamma> equivalence_preserved_by_iso_right \<psi>'.equivalence_in_bicategory_axioms by simp
hence "quasi_inverses f g"
using quasi_inverses_def by blast
thus ?thesis
using equivalence_mapI quasi_inverses_symmetric by blast
qed
lemma left_adjoint_to_equivalence_is_equivalence:
assumes "equivalence_map f" and "adjoint_pair g f"
shows "equivalence_map g"
proof -
obtain \<eta> \<epsilon> where gf: "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g f \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret gf: adjunction_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg g f \<eta> \<epsilon>
using gf by simp
obtain g' where 1: "quasi_inverses g' f"
using assms equivalence_mapE quasi_inverses_symmetric by blast
obtain \<phi> \<psi> where g'f: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g' f \<phi> \<psi>"
using assms 1 quasi_inverses_def by auto
interpret g'f: equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg g' f \<phi> \<psi>
using g'f by auto
obtain \<psi>' where \<psi>': "adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g' f \<phi> \<psi>'"
using assms 1 equivalence_refines_to_adjoint_equivalence [of g' f \<phi>]
g'f.antipar g'f.unit_in_hom g'f.unit_is_iso quasi_inverses_def
equivalence_map_def
by auto
interpret \<psi>': adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg g' f \<phi> \<psi>'
using \<psi>' by simp
have 1: "g \<cong> g'"
using gf.adjunction_in_bicategory_axioms \<psi>'.adjunction_in_bicategory_axioms
right_adjoint_determines_left_up_to_iso adjoint_pair_def
by blast
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : g' \<Rightarrow> g\<guillemotright> \<and> iso \<gamma>"
using 1 isomorphic_def isomorphic_symmetric by metis
have "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g f ((f \<star> \<gamma>) \<cdot> \<phi>) (\<psi>' \<cdot> (inv \<gamma> \<star> f))"
using \<gamma> equivalence_preserved_by_iso_left \<psi>'.equivalence_in_bicategory_axioms by simp
hence "quasi_inverses g f"
using quasi_inverses_def by auto
thus ?thesis
using quasi_inverses_symmetric quasi_inverses_def equivalence_map_def by blast
qed
lemma quasi_inverses_are_adjoint_pair:
assumes "quasi_inverses f g"
shows "adjoint_pair f g"
proof -
obtain \<eta> \<epsilon> where \<eta>\<epsilon>: "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms quasi_inverses_def by auto
interpret \<eta>\<epsilon>: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using \<eta>\<epsilon> by auto
obtain \<epsilon>' where \<eta>\<epsilon>': "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'"
using \<eta>\<epsilon> equivalence_map_def \<eta>\<epsilon>.antipar \<eta>\<epsilon>.unit_in_hom \<eta>\<epsilon>.unit_is_iso
\<eta>\<epsilon>.ide_right equivalence_refines_to_adjoint_equivalence [of f g \<eta>]
by force
interpret \<eta>\<epsilon>': adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using \<eta>\<epsilon>' by auto
show ?thesis
using \<eta>\<epsilon>' adjoint_pair_def \<eta>\<epsilon>'.adjunction_in_bicategory_axioms by auto
qed
lemma quasi_inverses_isomorphic_right:
assumes "quasi_inverses f g"
shows "quasi_inverses f g' \<longleftrightarrow> g \<cong> g'"
proof
show "g \<cong> g' \<Longrightarrow> quasi_inverses f g'"
using assms quasi_inverses_def isomorphic_def equivalence_preserved_by_iso_right
by metis
assume g': "quasi_inverses f g'"
show "g \<cong> g'"
using assms g' quasi_inverses_are_adjoint_pair left_adjoint_determines_right_up_to_iso
by blast
qed
lemma quasi_inverses_isomorphic_left:
assumes "quasi_inverses f g"
shows "quasi_inverses f' g \<longleftrightarrow> f \<cong> f'"
proof
show "f \<cong> f' \<Longrightarrow> quasi_inverses f' g"
using assms quasi_inverses_def isomorphic_def equivalence_preserved_by_iso_left
by metis
assume f': "quasi_inverses f' g"
show "f \<cong> f'"
using assms f' quasi_inverses_are_adjoint_pair right_adjoint_determines_left_up_to_iso
by blast
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,1445 +1,1445 @@
(* Title: InternalEquivalence
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Internal Equivalences"
theory InternalEquivalence
imports Bicategory
begin
text \<open>
An \emph{internal equivalence} in a bicategory consists of antiparallel 1-cells \<open>f\<close> and \<open>g\<close>
together with invertible 2-cells \<open>\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>\<close> and \<open>\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> src g\<guillemotright>\<close>.
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.
\<close>
subsection "Definition of Equivalence"
text \<open>
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.
\<close>
locale adjunction_data_in_bicategory =
bicategory +
fixes f :: 'a
and g :: 'a
and \<eta> :: 'a
and \<epsilon> :: 'a
assumes ide_left [simp]: "ide f"
and ide_right [simp]: "ide g"
and unit_in_vhom: "\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>"
and counit_in_vhom: "\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> src g\<guillemotright>"
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 "\<guillemotleft>\<epsilon> : trg f \<rightarrow> trg f\<guillemotright>" and "\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> trg f\<guillemotright>"
using counit_in_vhom vconn_implies_hpar antipar by auto
lemma unit_in_hom [intro]:
shows "\<guillemotleft>\<eta> : src f \<rightarrow> src f\<guillemotright>" and "\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>"
using unit_in_vhom vconn_implies_hpar antipar by auto
lemma unit_simps [simp]:
shows "arr \<eta>" and "dom \<eta> = src f" and "cod \<eta> = g \<star> f"
and "src \<eta> = src f" and "trg \<eta> = src f"
using unit_in_hom antipar by auto
lemma counit_simps [simp]:
shows "arr \<epsilon>" and "dom \<epsilon> = f \<star> g" and "cod \<epsilon> = trg f"
and "src \<epsilon> = trg f" and "trg \<epsilon> = trg f"
using counit_in_hom antipar by auto
text \<open>
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.
\<close>
lemma triangle_in_hom:
shows "\<guillemotleft>(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) : f \<star> src f \<Rightarrow> trg f \<star> f\<guillemotright>"
and "\<guillemotleft>(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) : trg g \<star> g \<Rightarrow> g \<star> src g\<guillemotright>"
and "\<guillemotleft>\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] : f \<Rightarrow> f\<guillemotright>"
and "\<guillemotleft>\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] : g \<Rightarrow> g\<guillemotright>"
using antipar by auto
lemma triangle_equiv_form:
shows "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<longleftrightarrow>
\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] = f"
and "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<longleftrightarrow>
\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] = g"
proof -
show "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<longleftrightarrow>
\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] = f"
proof
assume 1: "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
have "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] =
\<l>[f] \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) \<cdot> \<r>\<^sup>-\<^sup>1[f]"
using comp_assoc by simp
also have "... = \<l>[f] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<r>\<^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 "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] = f"
by simp
next
assume 2: "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] = f"
have "\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] = \<l>\<^sup>-\<^sup>1[f] \<cdot> f \<cdot> \<r>[f]"
using comp_cod_arr by simp
also have "... = (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<l>[f]) \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) \<cdot> (\<r>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
using 2 comp_assoc by (metis (no_types, lifting))
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)"
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 "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
by simp
qed
show "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<longleftrightarrow>
\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] = g"
proof
assume 1: "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
have "\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] =
\<r>[g] \<cdot> ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot> \<l>\<^sup>-\<^sup>1[g]"
using comp_assoc by simp
also have "... = \<r>[g] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<l>\<^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 "\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] = g"
by simp
next
assume 2: "\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] = g"
have "\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] = \<r>\<^sup>-\<^sup>1[g] \<cdot> g \<cdot> \<l>[g]"
using comp_cod_arr by simp
also have "... = \<r>\<^sup>-\<^sup>1[g] \<cdot> (\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g]) \<cdot> \<l>[g]"
using 2 by simp
also have "... = (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<r>[g]) \<cdot> ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot> (\<l>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g])"
using comp_assoc by simp
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> 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 \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
by simp
qed
qed
end
locale equivalence_in_bicategory =
adjunction_data_in_bicategory +
assumes unit_is_iso [simp]: "iso \<eta>"
and counit_is_iso [simp]: "iso \<epsilon>"
begin
lemma dual_equivalence:
shows "equivalence_in_bicategory V H \<a> \<i> src trg g f (inv \<epsilon>) (inv \<eta>)"
using antipar by unfold_locales auto
end
abbreviation (in bicategory) internal_equivalence
where "internal_equivalence f g \<phi> \<psi> \<equiv> equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>"
subsection "Quasi-Inverses and Equivalence Maps"
text \<open>
Antiparallel 1-cells \<open>f\<close> and \<open>g\<close> 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.
\<close>
context bicategory
begin
definition quasi_inverses
where "quasi_inverses f g \<equiv> \<exists>\<phi> \<psi>. internal_equivalence f g \<phi> \<psi>"
lemma quasi_inversesI:
assumes "ide f" and "ide g"
and "src f \<cong> g \<star> f" and "f \<star> g \<cong> 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 \<phi> where \<phi>: "\<guillemotleft>\<phi> : src f \<Rightarrow> g \<star> f\<guillemotright> \<and> iso \<phi>"
using assms isomorphic_def by auto
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : f \<star> g \<Rightarrow> trg f\<guillemotright> \<and> iso \<psi>"
using assms isomorphic_def by auto
have "equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>"
using assms 1 \<phi> \<psi> by unfold_locales auto
thus "\<exists>\<phi> \<psi>. internal_equivalence f g \<phi> \<psi>" by auto
qed
lemma quasi_inversesE:
assumes "quasi_inverses f g"
and "\<lbrakk>ide f; ide g; src f \<cong> g \<star> f; f \<star> g \<cong> trg f\<rbrakk> \<Longrightarrow> T"
shows T
proof -
obtain \<phi> \<psi> where \<phi>\<psi>: "internal_equivalence f g \<phi> \<psi>"
using assms quasi_inverses_def by auto
interpret \<phi>\<psi>: equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>
using \<phi>\<psi> by simp
have "ide f \<and> ide g"
by simp
moreover have "src f \<cong> g \<star> f"
using isomorphic_def \<phi>\<psi>.unit_in_hom by auto
moreover have "f \<star> g \<cong> trg f"
using isomorphic_def \<phi>\<psi>.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 \<phi> \<psi> where \<phi>\<psi>: "internal_equivalence f g \<phi> \<psi>"
using assms quasi_inverses_def by auto
interpret \<phi>\<psi>: equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>
using \<phi>\<psi> by simp
obtain \<phi>' \<psi>' where \<phi>'\<psi>': "internal_equivalence f g' \<phi>' \<psi>'"
using assms quasi_inverses_def by auto
interpret \<phi>'\<psi>': equivalence_in_bicategory V H \<a> \<i> src trg f g' \<phi>' \<psi>'
using \<phi>'\<psi>' by simp
have "\<guillemotleft>\<r>[g'] \<cdot> (g' \<star> \<psi>) \<cdot> \<a>[g', f, g] \<cdot> (\<phi>' \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] : g \<Rightarrow> g'\<guillemotright>"
using \<phi>\<psi>.unit_in_hom \<phi>\<psi>.unit_is_iso \<phi>\<psi>.antipar \<phi>'\<psi>'.antipar
by (intro comp_in_homI' hseqI') auto
moreover have "iso (\<r>[g'] \<cdot> (g' \<star> \<psi>) \<cdot> \<a>[g', f, g] \<cdot> (\<phi>' \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g])"
using \<phi>\<psi>.unit_in_hom \<phi>\<psi>.unit_is_iso \<phi>\<psi>.antipar \<phi>'\<psi>'.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 \<equiv> \<exists>g \<eta> \<epsilon>. equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
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 \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>" and "\<guillemotleft>\<psi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<psi>"
shows "internal_equivalence f' g' ((g' \<star> \<phi>) \<cdot> (\<psi> \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g) \<cdot> (f' \<star> inv \<psi>))"
proof -
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
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: "\<guillemotleft>(g' \<star> \<phi>) \<cdot> (\<psi> \<star> f) \<cdot> \<eta> : src f' \<Rightarrow> g' \<star> f'\<guillemotright>"
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' \<star> \<phi>) \<cdot> (\<psi> \<star> f) \<cdot> \<eta>)"
using assms 1 g' vconn_implies_hpar(3) E.antipar(2) iso_hcomp
by (intro isos_compose) auto
show 1: "\<guillemotleft>\<epsilon> \<cdot> (inv \<phi> \<star> g) \<cdot> (f' \<star> inv \<psi>) : f' \<star> g' \<Rightarrow> src g'\<guillemotright>"
using assms f' ide_in_hom(2) vconn_implies_hpar(3-4) E.antipar(1-2)
by (intro comp_in_homI) auto
show "iso (\<epsilon> \<cdot> (inv \<phi> \<star> g) \<cdot> (f' \<star> inv \<psi>))"
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 \<cong> f'"
shows "equivalence_map f'"
proof -
obtain g \<eta> \<epsilon> where E: "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms equivalence_map_def by auto
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using E by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright> \<and> iso \<phi>"
using assms isomorphic_def isomorphic_symmetric by blast
have "equivalence_in_bicategory V H \<a> \<i> src trg f' g
((g \<star> \<phi>) \<cdot> (g \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g) \<cdot> (f' \<star> inv g))"
using E \<phi> equivalence_respects_iso [of f g \<eta> \<epsilon> \<phi> 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 \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "equivalence_in_bicategory V H \<a> \<i> src trg f g' ((\<phi> \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (f \<star> inv \<phi>))"
proof
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by auto
show "ide f" by simp
show "ide g'"
using assms(2) isomorphic_def by auto
show "\<guillemotleft>(\<phi> \<star> f) \<cdot> \<eta> : src f \<Rightarrow> g' \<star> f\<guillemotright>"
using assms E.antipar(2) E.ide_left by blast
show "\<guillemotleft>\<epsilon> \<cdot> (f \<star> inv \<phi>) : f \<star> g' \<Rightarrow> src g'\<guillemotright>"
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 ((\<phi> \<star> f) \<cdot> \<eta>)"
using assms E.antipar isos_compose by auto
show "iso (\<epsilon> \<cdot> (f \<star> inv \<phi>))"
using assms E.antipar isos_compose by auto
qed
lemma equivalence_preserved_by_iso_left:
assumes "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "equivalence_in_bicategory V H \<a> \<i> src trg f' g ((g \<star> \<phi>) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g))"
proof
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by auto
show "ide f'"
using assms by auto
show "ide g"
by simp
show "\<guillemotleft>(g \<star> \<phi>) \<cdot> \<eta> : src f' \<Rightarrow> g \<star> f'\<guillemotright>"
using assms E.unit_in_hom E.antipar by auto
show "\<guillemotleft>\<epsilon> \<cdot> (inv \<phi> \<star> g) : f' \<star> g \<Rightarrow> src g\<guillemotright>"
using assms E.counit_in_hom E.antipar ide_in_hom(2) vconn_implies_hpar(3) by auto
show "iso ((g \<star> \<phi>) \<cdot> \<eta>)"
using assms E.antipar isos_compose by auto
show "iso (\<epsilon> \<cdot> (inv \<phi> \<star> 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 "\<lambda>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 \<phi> \<psi> where \<phi>\<psi>: "equivalence_in_bicategory V H \<a> \<i> src trg f f\<^sup>~ \<phi> \<psi>"
using assms equivalence_map_def quasi_inverses_some_quasi_inverse quasi_inverses_def
by blast
interpret \<phi>\<psi>: equivalence_in_bicategory V H \<a> \<i> src trg f \<open>f\<^sup>~\<close> \<phi> \<psi>
using \<phi>\<psi> by simp
show "src f\<^sup>~ = trg f"
using \<phi>\<psi>.antipar by simp
show "trg f\<^sup>~ = src f"
using \<phi>\<psi>.antipar by simp
qed
lemma quasi_inverse_in_hom [intro]:
assumes "equivalence_map f"
shows "\<guillemotleft>f\<^sup>~ : trg f \<rightarrow> src f\<guillemotright>"
and "\<guillemotleft>f\<^sup>~ : f\<^sup>~ \<Rightarrow> f\<^sup>~\<guillemotright>"
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
apply auto
by (meson equivalence_mapI)
lemma quasi_inverse_quasi_inverse:
assumes "equivalence_map f"
shows "(f\<^sup>~)\<^sup>~ \<cong> 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>~ \<star> f \<cong> src f" and "f \<star> f\<^sup>~ \<cong> trg f"
proof -
obtain \<phi> \<psi> where \<phi>\<psi>: "equivalence_in_bicategory V H \<a> \<i> src trg f f\<^sup>~ \<phi> \<psi>"
using assms equivalence_map_def quasi_inverses_some_quasi_inverse
quasi_inverses_def
by blast
interpret \<phi>\<psi>: equivalence_in_bicategory V H \<a> \<i> src trg f \<open>f\<^sup>~\<close> \<phi> \<psi>
using \<phi>\<psi> by simp
show "f\<^sup>~ \<star> f \<cong> src f"
using quasi_inverses_some_quasi_inverse quasi_inverses_def
\<phi>\<psi>.unit_in_hom \<phi>\<psi>.unit_is_iso isomorphic_def
by blast
show "f \<star> f\<^sup>~ \<cong> trg f"
using quasi_inverses_some_quasi_inverse quasi_inverses_def
\<phi>\<psi>.counit_in_hom \<phi>\<psi>.counit_is_iso isomorphic_def
by blast
qed
lemma quasi_inverse_transpose:
assumes "ide f" and "ide g" and "ide h" and "f \<star> g \<cong> h"
shows "equivalence_map g \<Longrightarrow> f \<cong> h \<star> g\<^sup>~"
and "equivalence_map f \<Longrightarrow> g \<cong> f\<^sup>~ \<star> h"
proof -
have 1: "src f = trg g"
using assms equivalence_map_is_ide by fastforce
show "equivalence_map g \<Longrightarrow> f \<cong> h \<star> g\<^sup>~"
proof -
assume g: "equivalence_map g"
have 2: "ide g\<^sup>~"
using g by simp
have "f \<cong> f \<star> src f"
using assms isomorphic_unit_right isomorphic_symmetric by blast
also have "... \<cong> f \<star> trg g"
using assms 1 isomorphic_reflexive by auto
also have "... \<cong> f \<star> g \<star> g\<^sup>~"
using assms g 1 comp_quasi_inverse(2) isomorphic_symmetric hcomp_ide_isomorphic
by simp
also have "... \<cong> (f \<star> g) \<star> g\<^sup>~"
using assms g 1 2 assoc'_in_hom [of f g "g\<^sup>~"] iso_assoc' isomorphic_def by auto
also have "... \<cong> h \<star> g\<^sup>~"
using assms g 1 2
by (simp add: hcomp_isomorphic_ide)
finally show ?thesis by blast
qed
show "equivalence_map f \<Longrightarrow> g \<cong> f\<^sup>~ \<star> h"
proof -
assume f: "equivalence_map f"
have 2: "ide f\<^sup>~"
using f by simp
have "g \<cong> src f \<star> g"
using assms 1 isomorphic_unit_left isomorphic_symmetric by metis
also have "... \<cong> (f\<^sup>~ \<star> f) \<star> g"
using assms f 1 comp_quasi_inverse(1) [of f] isomorphic_symmetric
hcomp_isomorphic_ide
by simp
also have "... \<cong> f\<^sup>~ \<star> f \<star> g"
using assms f 1 assoc_in_hom [of "f\<^sup>~" f g] iso_assoc isomorphic_def by auto
also have "... \<cong> f\<^sup>~ \<star> 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 \<a> \<i> src trg +
fg: equivalence_in_bicategory V H \<a> \<i> src trg f g \<zeta> \<xi> +
hk: equivalence_in_bicategory V H \<a> \<i> src trg h k \<sigma> \<tau>
for V :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and f :: "'a"
and g :: "'a"
and \<zeta> :: "'a"
and \<xi> :: "'a"
and h :: "'a"
and k :: "'a"
and \<sigma> :: "'a"
and \<tau> :: "'a" +
assumes composable: "src h = trg f"
begin
abbreviation \<eta>
where "\<eta> \<equiv> \<a>\<^sup>-\<^sup>1[g, k, h \<star> f] \<cdot> (g \<star> \<a>[k, h, f]) \<cdot> (g \<star> \<sigma> \<star> f) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[f]) \<cdot> \<zeta>"
abbreviation \<epsilon>
where "\<epsilon> \<equiv> \<tau> \<cdot> (h \<star> \<l>[k]) \<cdot> (h \<star> \<xi> \<star> k) \<cdot> (h \<star> \<a>\<^sup>-\<^sup>1[f, g, k]) \<cdot> \<a>[h, f, g \<star> k]"
interpretation adjunction_data_in_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
proof
show "ide (h \<star> f)"
using composable by simp
show "ide (g \<star> k)"
using fg.antipar hk.antipar composable by simp
show "\<guillemotleft>\<eta> : src (h \<star> f) \<Rightarrow> (g \<star> k) \<star> h \<star> f\<guillemotright>"
using fg.antipar hk.antipar composable by fastforce
show "\<guillemotleft>\<epsilon> : (h \<star> f) \<star> g \<star> k \<Rightarrow> src (g \<star> k)\<guillemotright>"
using fg.antipar hk.antipar composable by fastforce
qed
interpretation equivalence_in_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
proof
show "iso \<eta>"
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 \<epsilon>"
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 \<a> \<i> src trg (h \<star> f) (g \<star> k) \<eta> \<epsilon>"
..
sublocale equivalence_in_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
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' \<star> f)"
proof -
obtain g \<phi> \<psi> where fg: "equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>"
using assms(1) equivalence_map_def by auto
interpret fg: equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>
using fg by auto
obtain g' \<phi>' \<psi>' where f'g': "equivalence_in_bicategory V H \<a> \<i> src trg f' g' \<phi>' \<psi>'"
using assms(2) equivalence_map_def by auto
interpret f'g': equivalence_in_bicategory V H \<a> \<i> src trg f' g' \<phi>' \<psi>'
using f'g' by auto
interpret composite_equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi> f' g' \<phi>' \<psi>'
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 \<star> f')"
and "quasi_inverses f g" and "quasi_inverses f' g'"
shows "quasi_inverses (f \<star> f') (g' \<star> g)"
proof -
obtain \<phi> \<psi> where g: "internal_equivalence f g \<phi> \<psi>"
using assms quasi_inverses_def by auto
interpret g: equivalence_in_bicategory V H \<a> \<i> src trg f g \<phi> \<psi>
using g by simp
obtain \<phi>' \<psi>' where g': "internal_equivalence f' g' \<phi>' \<psi>'"
using assms quasi_inverses_def by auto
interpret g': equivalence_in_bicategory V H \<a> \<i> src trg f' g' \<phi>' \<psi>'
using g' by simp
interpret gg': composite_equivalence_in_bicategory V H \<a> \<i> src trg f' g' \<phi>' \<psi>' f g \<phi> \<psi>
using assms equivalence_map_is_ide [of "f \<star> 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 \<star> f')"
shows "(f \<star> f')\<^sup>~ \<cong> f'\<^sup>~ \<star> 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 \<cong> f'"
shows "f\<^sup>~ \<cong> f'\<^sup>~"
proof -
have hpar: "src f = src f' \<and> trg f = trg f'"
using assms isomorphic_implies_hpar by simp
have "f\<^sup>~ \<cong> f\<^sup>~ \<star> trg f"
using isomorphic_unit_right
by (metis assms(1) isomorphic_symmetric quasi_inverse_simps(2-3))
also have "... \<cong> f\<^sup>~ \<star> f' \<star> f'\<^sup>~"
proof -
have "trg f \<cong> f' \<star> 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 "... \<cong> (f\<^sup>~ \<star> f') \<star> f'\<^sup>~"
using assms hcomp_assoc_isomorphic hpar isomorphic_implies_ide(2) isomorphic_symmetric
by auto
also have "... \<cong> (f\<^sup>~ \<star> f) \<star> f'\<^sup>~"
proof -
have "f\<^sup>~ \<star> f' \<cong> f\<^sup>~ \<star> 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 "... \<cong> src f \<star> f'\<^sup>~"
proof -
have "f\<^sup>~ \<star> f \<cong> src f"
using assms comp_quasi_inverse by simp
thus ?thesis
using assms hcomp_isomorphic_ide isomorphic_implies_hpar by simp
qed
also have "... \<cong> 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 \<equiv> \<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> 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: "\<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> equivalence_map f"
using assms equivalent_objects_def by auto
obtain g where g: "\<guillemotleft>g : b \<rightarrow> c\<guillemotright> \<and> equivalence_map g"
using assms equivalent_objects_def by auto
have "\<guillemotleft>g \<star> f : a \<rightarrow> c\<guillemotright> \<and> equivalence_map (g \<star> 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 \<open>
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.
\<close>
locale two_equivalences_in_bicategory =
bicategory V H \<a> \<i> src trg +
e\<^sub>0: equivalence_in_bicategory V H \<a> \<i> src trg e\<^sub>0 d\<^sub>0 \<eta>\<^sub>0 \<epsilon>\<^sub>0 +
e\<^sub>1: equivalence_in_bicategory V H \<a> \<i> src trg e\<^sub>1 d\<^sub>1 \<eta>\<^sub>1 \<epsilon>\<^sub>1
for V :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and e\<^sub>0 :: "'a"
and d\<^sub>0 :: "'a"
and \<eta>\<^sub>0 :: "'a"
and \<epsilon>\<^sub>0 :: "'a"
and e\<^sub>1 :: "'a"
and d\<^sub>1 :: "'a"
and \<eta>\<^sub>1 :: "'a"
and \<epsilon>\<^sub>1 :: "'a"
begin
interpretation hom: subcategory V \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>\<close>
using hhom_is_subcategory by simp
(* TODO: The preceding interpretation somehow brings in unwanted notation. *)
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
interpretation hom': subcategory V \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>\<close>
using hhom_is_subcategory by simp
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
abbreviation (input) F
where "F \<equiv> \<lambda>\<mu>. e\<^sub>1 \<star> \<mu> \<star> d\<^sub>0"
interpretation F: "functor" hom.comp hom'.comp F
proof
show 1: "\<And>f. hom.arr f \<Longrightarrow> hom'.arr (e\<^sub>1 \<star> f \<star> d\<^sub>0)"
- using hom.arr_char hom'.arr_char in_hhom_def e\<^sub>0.antipar(1-2) by simp
+ using hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C in_hhom_def e\<^sub>0.antipar(1-2) by simp
show "\<And>f. \<not> hom.arr f \<Longrightarrow> e\<^sub>1 \<star> f \<star> d\<^sub>0 = hom'.null"
- using 1 hom.arr_char hom'.null_char in_hhom_def
+ using 1 hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.null_char in_hhom_def
by (metis e\<^sub>0.antipar(1) hseqE hseq_char' hcomp_simps(2))
show "\<And>f. hom.arr f \<Longrightarrow> hom'.dom (e\<^sub>1 \<star> f \<star> d\<^sub>0) = e\<^sub>1 \<star> hom.dom f \<star> d\<^sub>0"
- using hom.arr_char hom.dom_char hom'.arr_char hom'.dom_char
+ using hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.dom_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.dom_char\<^sub>S\<^sub>b\<^sub>C
by (metis 1 hcomp_simps(3) e\<^sub>0.ide_right e\<^sub>1.ide_left hom'.inclusion hseq_char ide_char)
show "\<And>f. hom.arr f \<Longrightarrow> hom'.cod (e\<^sub>1 \<star> f \<star> d\<^sub>0) = e\<^sub>1 \<star> hom.cod f \<star> d\<^sub>0"
- using hom.arr_char hom.cod_char hom'.arr_char hom'.cod_char
+ using hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.cod_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis 1 hcomp_simps(4) e\<^sub>0.ide_right e\<^sub>1.ide_left hom'.inclusion hseq_char ide_char)
show "\<And>g f. hom.seq g f \<Longrightarrow>
e\<^sub>1 \<star> hom.comp g f \<star> d\<^sub>0 = hom'.comp (e\<^sub>1 \<star> g \<star> d\<^sub>0) (e\<^sub>1 \<star> f \<star> d\<^sub>0)"
- using 1 hom.seq_char hom.arr_char hom.comp_char hom'.arr_char hom'.comp_char
+ using 1 hom.seq_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.comp_char hom'.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<equiv> \<lambda>\<mu>'. d\<^sub>1 \<star> \<mu>' \<star> e\<^sub>0"
interpretation G: "functor" hom'.comp hom.comp G
proof
show 1: "\<And>f. hom'.arr f \<Longrightarrow> hom.arr (d\<^sub>1 \<star> f \<star> e\<^sub>0)"
- using hom.arr_char hom'.arr_char in_hhom_def e\<^sub>1.antipar(1) e\<^sub>1.antipar(2)
+ using hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C in_hhom_def e\<^sub>1.antipar(1) e\<^sub>1.antipar(2)
by simp
show "\<And>f. \<not> hom'.arr f \<Longrightarrow> d\<^sub>1 \<star> f \<star> 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))
+ using 1 hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.null_char in_hhom_def
+ by (metis e\<^sub>1.antipar(2) hom'.arrI\<^sub>S\<^sub>b\<^sub>C hom.null_char hseqE hseq_char' hcomp_simps(2))
show "\<And>f. hom'.arr f \<Longrightarrow> hom.dom (d\<^sub>1 \<star> f \<star> e\<^sub>0) = d\<^sub>1 \<star> hom'.dom f \<star> e\<^sub>0"
- using 1 hom.arr_char hom.dom_char hom'.arr_char hom'.dom_char
+ using 1 hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.dom_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.dom_char\<^sub>S\<^sub>b\<^sub>C
by (metis hcomp_simps(3) e\<^sub>0.ide_left e\<^sub>1.ide_right hom.inclusion hseq_char ide_char)
show "\<And>f. hom'.arr f \<Longrightarrow> hom.cod (d\<^sub>1 \<star> f \<star> e\<^sub>0) = d\<^sub>1 \<star> hom'.cod f \<star> e\<^sub>0"
- using 1 hom.arr_char hom.cod_char hom'.arr_char hom'.cod_char
+ using 1 hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.cod_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis hcomp_simps(4) e\<^sub>0.ide_left e\<^sub>1.ide_right hom.inclusion hseq_char ide_char)
show "\<And>g f. hom'.seq g f \<Longrightarrow>
d\<^sub>1 \<star> hom'.comp g f \<star> e\<^sub>0 = hom.comp (d\<^sub>1 \<star> g \<star> e\<^sub>0) (d\<^sub>1 \<star> f \<star> e\<^sub>0)"
- using 1 hom'.seq_char hom'.arr_char hom'.comp_char hom.arr_char hom.comp_char
+ using 1 hom'.seq_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.comp_char hom.arr_char\<^sub>S\<^sub>b\<^sub>C 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) \<phi>\<^sub>0
where "\<phi>\<^sub>0 f \<equiv> (d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, f \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (f \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[f, d\<^sub>0, e\<^sub>0]) \<cdot> (\<eta>\<^sub>1 \<star> f \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[f]"
lemma \<phi>\<^sub>0_in_hom:
assumes "\<guillemotleft>f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>" and "ide f"
shows "\<guillemotleft>\<phi>\<^sub>0 f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
and "\<guillemotleft>\<phi>\<^sub>0 f : f \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> f \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
proof -
show "\<guillemotleft>\<phi>\<^sub>0 f : f \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> f \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
using assms e\<^sub>0.antipar e\<^sub>1.antipar by fastforce
thus "\<guillemotleft>\<phi>\<^sub>0 f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
using assms src_dom [of "\<phi>\<^sub>0 f"] trg_dom [of "\<phi>\<^sub>0 f"]
by (metis arrI dom_comp in_hhom_def runit'_simps(4) seqE)
qed
lemma iso_\<phi>\<^sub>0:
assumes "\<guillemotleft>f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>" and "ide f"
shows "iso (\<phi>\<^sub>0 f)"
using assms iso_lunit' iso_runit' e\<^sub>0.antipar e\<^sub>1.antipar
by (intro isos_compose) auto
interpretation \<phi>: transformation_by_components hom.comp hom.comp hom.map \<open>G o F\<close> \<phi>\<^sub>0
proof
fix f
assume f: "hom.ide f"
show "hom.in_hom (\<phi>\<^sub>0 f) (hom.map f) (GF.map f)"
- proof (unfold hom.in_hom_char, intro conjI)
+ proof (unfold hom.in_hom_char\<^sub>S\<^sub>b\<^sub>C, 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 (\<phi>\<^sub>0 f)"
- using f hom.ide_char hom.arr_char \<phi>\<^sub>0_in_hom by simp
+ using f hom.ide_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C \<phi>\<^sub>0_in_hom by simp
show "\<guillemotleft>\<phi>\<^sub>0 f : hom.map f \<Rightarrow> GF.map f\<guillemotright>"
- using f hom.ide_char hom.arr_char hom'.ide_char hom'.arr_char F.preserves_arr
+ using f hom.ide_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.ide_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C F.preserves_arr
\<phi>\<^sub>0_in_hom
by auto
qed
next
fix \<mu>
assume \<mu>: "hom.arr \<mu>"
show "hom.comp (\<phi>\<^sub>0 (hom.cod \<mu>)) (hom.map \<mu>) =
hom.comp (GF.map \<mu>) (\<phi>\<^sub>0 (hom.dom \<mu>))"
proof -
have "hom.comp (\<phi>\<^sub>0 (hom.cod \<mu>)) (hom.map \<mu>) = \<phi>\<^sub>0 (cod \<mu>) \<cdot> \<mu>"
proof -
have "hom.map \<mu> = \<mu>"
using \<mu> by simp
moreover have "\<phi>\<^sub>0 (hom.cod \<mu>) = \<phi>\<^sub>0 (cod \<mu>)"
- using \<mu> hom.arr_char hom.cod_char by simp
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.cod_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "hom.arr (\<phi>\<^sub>0 (cod \<mu>))"
- using \<mu> hom.arr_char \<phi>\<^sub>0_in_hom [of "cod \<mu>"]
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C \<phi>\<^sub>0_in_hom [of "cod \<mu>"]
by (metis hom.arr_cod hom.cod_simp ide_cod in_hhom_def)
moreover have "seq (\<phi>\<^sub>0 (cod \<mu>)) \<mu>"
proof
show 1: "\<guillemotleft>\<mu> : dom \<mu> \<Rightarrow> cod \<mu>\<guillemotright>"
- using \<mu> hom.arr_char by auto
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
show "\<guillemotleft>\<phi>\<^sub>0 (cod \<mu>) : cod \<mu> \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
- using \<mu> hom.arr_char \<phi>\<^sub>0_in_hom
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C \<phi>\<^sub>0_in_hom
by (metis 1 arrI hom.arr_cod hom.cod_simp ide_cod)
qed
ultimately show ?thesis
using \<mu> hom.comp_char by simp
qed
also have "... = (d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot> (\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0) \<cdot>
\<l>\<^sup>-\<^sup>1[cod \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[cod \<mu>] \<cdot> \<mu>"
- using \<mu> hom.arr_char comp_assoc by auto
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C comp_assoc by auto
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot> (\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0) \<cdot>
\<l>\<^sup>-\<^sup>1[cod \<mu> \<star> src e\<^sub>0] \<cdot> (\<mu> \<star> src e\<^sub>0)) \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
- using \<mu> hom.arr_char runit'_naturality [of \<mu>] comp_assoc by auto
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C runit'_naturality [of \<mu>] comp_assoc by auto
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot> (\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0) \<cdot>
(src e\<^sub>1 \<star> \<mu> \<star> src e\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0]) \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
- using \<mu> hom.arr_char lunit'_naturality [of "\<mu> \<star> src e\<^sub>0"] by force
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C lunit'_naturality [of "\<mu> \<star> src e\<^sub>0"] by force
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot> (\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0) \<cdot>
(src e\<^sub>1 \<star> \<mu> \<star> src e\<^sub>0)) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
using comp_assoc by simp
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot> ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<mu> \<star> d\<^sub>0 \<star> e\<^sub>0)) \<cdot>
(\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
proof -
have "(\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0) \<cdot> (src e\<^sub>1 \<star> \<mu> \<star> src e\<^sub>0) = \<eta>\<^sub>1 \<star> \<mu> \<star> \<eta>\<^sub>0"
- using \<mu> hom.arr_char comp_arr_dom comp_cod_arr
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C comp_arr_dom comp_cod_arr
interchange [of \<eta>\<^sub>1 "src e\<^sub>1" "cod \<mu> \<star> \<eta>\<^sub>0" "\<mu> \<star> src e\<^sub>0"]
interchange [of "cod \<mu>" \<mu> \<eta>\<^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 \<star> e\<^sub>1) \<star> \<mu> \<star> d\<^sub>0 \<star> e\<^sub>0) \<cdot> (\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0)"
proof -
have "\<eta>\<^sub>1 \<star> \<mu> \<star> \<eta>\<^sub>0 = (d\<^sub>1 \<star> e\<^sub>1) \<cdot> \<eta>\<^sub>1 \<star> \<mu> \<cdot> dom \<mu> \<star> (d\<^sub>0 \<star> e\<^sub>0) \<cdot> \<eta>\<^sub>0"
- using \<mu> hom.arr_char comp_arr_dom comp_cod_arr by auto
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C comp_arr_dom comp_cod_arr by auto
also have "... = (d\<^sub>1 \<star> e\<^sub>1) \<cdot> \<eta>\<^sub>1 \<star> (\<mu> \<star> d\<^sub>0 \<star> e\<^sub>0) \<cdot> (dom \<mu> \<star> \<eta>\<^sub>0)"
- using \<mu> hom.arr_char comp_cod_arr
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C comp_cod_arr
interchange [of \<mu> "dom \<mu>" "d\<^sub>0 \<star> e\<^sub>0" \<eta>\<^sub>0]
by auto
also have "... = ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<mu> \<star> d\<^sub>0 \<star> e\<^sub>0) \<cdot> (\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0)"
- using \<mu> hom.arr_char comp_arr_dom comp_cod_arr
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C comp_arr_dom comp_cod_arr
interchange [of "d\<^sub>1 \<star> e\<^sub>1" \<eta>\<^sub>1 "\<mu> \<star> d\<^sub>0 \<star> e\<^sub>0" "dom \<mu> \<star> \<eta>\<^sub>0"]
interchange [of \<mu> "dom \<mu>" "d\<^sub>0 \<star> e\<^sub>0" \<eta>\<^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 "(\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0) \<cdot> (src e\<^sub>1 \<star> \<mu> \<star> src e\<^sub>0) =
((d\<^sub>1 \<star> e\<^sub>1) \<star> \<mu> \<star> d\<^sub>0 \<star> e\<^sub>0) \<cdot> (\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0)"
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> (\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0) \<cdot> ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[dom \<mu>, d\<^sub>0, e\<^sub>0])) \<cdot>
(\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
- using \<mu> hom.arr_char e\<^sub>0.antipar e\<^sub>1.antipar assoc'_naturality [of \<mu> d\<^sub>0 e\<^sub>0]
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C e\<^sub>0.antipar e\<^sub>1.antipar assoc'_naturality [of \<mu> d\<^sub>0 e\<^sub>0]
whisker_left [of "d\<^sub>1 \<star> e\<^sub>1" "\<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0]" "\<mu> \<star> d\<^sub>0 \<star> e\<^sub>0"]
whisker_left [of "d\<^sub>1 \<star> e\<^sub>1" "(\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0" "\<a>\<^sup>-\<^sup>1[dom \<mu>, d\<^sub>0, e\<^sub>0]"]
by auto
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> \<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot>
((d\<^sub>1 \<star> e\<^sub>1) \<star> (\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0)) \<cdot> ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[dom \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot>
(\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
using comp_assoc by simp
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> (d\<^sub>1 \<star> e\<^sub>1 \<star> (\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0) \<cdot>
\<a>[d\<^sub>1, e\<^sub>1, (dom \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0]) \<cdot> ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[dom \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot>
(\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
- using \<mu> hom.arr_char e\<^sub>0.antipar e\<^sub>1.antipar
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C e\<^sub>0.antipar e\<^sub>1.antipar
assoc_naturality [of d\<^sub>1 e\<^sub>1 "(\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0"] hseqI
by auto
also have "... = ((d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]) \<cdot> (d\<^sub>1 \<star> e\<^sub>1 \<star> (\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0)) \<cdot>
\<a>[d\<^sub>1, e\<^sub>1, (dom \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot> ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[dom \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot>
(\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
using comp_assoc by simp
also have "... = ((d\<^sub>1 \<star> (e\<^sub>1 \<star> \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0) \<cdot> (d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, dom \<mu> \<star> d\<^sub>0, e\<^sub>0])) \<cdot>
\<a>[d\<^sub>1, e\<^sub>1, (dom \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] \<cdot> ((d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[dom \<mu>, d\<^sub>0, e\<^sub>0]) \<cdot>
(\<eta>\<^sub>1 \<star> dom \<mu> \<star> \<eta>\<^sub>0) \<cdot> \<l>\<^sup>-\<^sup>1[dom \<mu> \<star> src e\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[dom \<mu>]"
- using \<mu> hom.arr_char e\<^sub>0.antipar e\<^sub>1.antipar
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C e\<^sub>0.antipar e\<^sub>1.antipar
assoc'_naturality [of e\<^sub>1 "\<mu> \<star> d\<^sub>0" e\<^sub>0]
whisker_left [of d\<^sub>1 "\<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0]" "e\<^sub>1 \<star> (\<mu> \<star> d\<^sub>0) \<star> e\<^sub>0"]
whisker_left [of d\<^sub>1 "(e\<^sub>1 \<star> \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0" "\<a>\<^sup>-\<^sup>1[e\<^sub>1, dom \<mu> \<star> d\<^sub>0, e\<^sub>0]"]
by auto
also have "... = hom.comp (GF.map \<mu>) (\<phi>\<^sub>0 (hom.dom \<mu>))"
proof -
have "hom.arr (GF.map \<mu>)"
using \<mu> by blast
moreover have "hom.arr (\<phi>\<^sub>0 (hom.dom \<mu>))"
- using \<mu> hom.arr_char hom.in_hom_char \<phi>\<^sub>0_in_hom(1) hom.dom_closed hom.dom_simp
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.in_hom_char\<^sub>S\<^sub>b\<^sub>C \<phi>\<^sub>0_in_hom(1) hom.dom_closed hom.dom_simp
hom.inclusion ide_dom
by metis
moreover have "seq (GF.map \<mu>) (\<phi>\<^sub>0 (hom.dom \<mu>))"
proof
show "\<guillemotleft>\<phi>\<^sub>0 (hom.dom \<mu>) : dom \<mu> \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> dom \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
- using \<mu> hom.arr_char hom.dom_char hom.in_hom_char \<phi>\<^sub>0_in_hom hom.dom_closed
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.dom_char\<^sub>S\<^sub>b\<^sub>C hom.in_hom_char\<^sub>S\<^sub>b\<^sub>C \<phi>\<^sub>0_in_hom hom.dom_closed
hom.dom_simp hom.inclusion ide_dom
by metis
show "\<guillemotleft>GF.map \<mu> : d\<^sub>1 \<star> (e\<^sub>1 \<star> dom \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0 \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
- using \<mu> hom.arr_char hom'.arr_char F.preserves_arr
+ using \<mu> hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu> hom.comp_char comp_assoc hom.dom_simp by auto
qed
finally show ?thesis by blast
qed
qed
lemma transformation_by_components_\<phi>\<^sub>0:
shows "transformation_by_components hom.comp hom.comp hom.map (G o F) \<phi>\<^sub>0"
..
interpretation \<phi>: natural_isomorphism hom.comp hom.comp hom.map \<open>G o F\<close> \<phi>.map
proof
fix f
assume "hom.ide f"
hence f: "ide f \<and> \<guillemotleft>f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
- using hom.ide_char hom.arr_char by simp
+ using hom.ide_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show "hom.iso (\<phi>.map f)"
- proof (unfold hom.iso_char hom.arr_char, intro conjI)
+ proof (unfold hom.iso_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C, intro conjI)
show "iso (\<phi>.map f)"
- using f iso_\<phi>\<^sub>0 \<phi>.map_simp_ide hom.ide_char hom.arr_char by simp
+ using f iso_\<phi>\<^sub>0 \<phi>.map_simp_ide hom.ide_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover show "\<guillemotleft>\<phi>.map f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
- using f hom.ide_char hom.arr_char by blast
+ using f hom.ide_char hom.arr_char\<^sub>S\<^sub>b\<^sub>C by blast
ultimately show "\<guillemotleft>inv (\<phi>.map f) : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
by auto
qed
qed
lemma natural_isomorphism_\<phi>:
shows "natural_isomorphism hom.comp hom.comp hom.map (G o F) \<phi>.map"
..
definition \<phi>
where "\<phi> \<equiv> \<phi>.map"
lemma \<phi>_ide_simp:
assumes "\<guillemotleft>f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>" and "ide f"
shows "\<phi> f = \<phi>\<^sub>0 f"
unfolding \<phi>_def
- using assms hom.ide_char hom.arr_char by simp
+ using assms hom.ide_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
lemma \<phi>_components_are_iso:
assumes "\<guillemotleft>f : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>" and "ide f"
shows "iso (\<phi> f)"
- using assms \<phi>_def \<phi>.components_are_iso hom.ide_char hom.arr_char hom.iso_char
+ using assms \<phi>_def \<phi>.components_are_iso hom.ide_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.iso_char\<^sub>S\<^sub>b\<^sub>C
by simp
lemma \<phi>_eq:
shows "\<phi> = (\<lambda>\<mu>. if \<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright> then \<phi>\<^sub>0 (cod \<mu>) \<cdot> \<mu> else null)"
proof
fix \<mu>
have "\<not> \<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright> \<Longrightarrow> \<phi>.map \<mu> = null"
- using hom.comp_char hom.null_char hom.arr_char
+ using hom.comp_char hom.null_char hom.arr_char\<^sub>S\<^sub>b\<^sub>C
by (simp add: \<phi>.is_extensional)
moreover have "\<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright> \<Longrightarrow> \<phi>.map \<mu> = \<phi>\<^sub>0 (cod \<mu>) \<cdot> \<mu>"
unfolding \<phi>.map_def
apply auto
- using hom.comp_char hom.arr_char hom.dom_simp hom.cod_simp
+ using hom.comp_char hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.dom_simp hom.cod_simp
apply simp
by (metis (no_types, lifting) \<phi>\<^sub>0_in_hom(1) hom.cod_closed hom.inclusion ide_cod local.ext)
ultimately show "\<phi> \<mu> = (if \<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright> then \<phi>\<^sub>0 (cod \<mu>) \<cdot> \<mu> else null)"
unfolding \<phi>_def by auto
qed
lemma \<phi>_in_hom [intro]:
assumes "\<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
shows "\<guillemotleft>\<phi> \<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
and "\<guillemotleft>\<phi> \<mu> : dom \<mu> \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
proof -
show "\<guillemotleft>\<phi> \<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
- using assms \<phi>_eq \<phi>_def hom.arr_char \<phi>.preserves_reflects_arr by presburger
+ using assms \<phi>_eq \<phi>_def hom.arr_char\<^sub>S\<^sub>b\<^sub>C \<phi>.preserves_reflects_arr by presburger
show "\<guillemotleft>\<phi> \<mu> : dom \<mu> \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
unfolding \<phi>_eq
using assms apply simp
apply (intro comp_in_homI)
apply auto
proof -
show "\<guillemotleft>\<r>\<^sup>-\<^sup>1[cod \<mu>] : cod \<mu> \<Rightarrow> cod \<mu> \<star> src e\<^sub>0\<guillemotright>"
using assms by auto
show "\<guillemotleft>\<l>\<^sup>-\<^sup>1[cod \<mu> \<star> src e\<^sub>0] : cod \<mu> \<star> src e\<^sub>0 \<Rightarrow> src e\<^sub>1 \<star> cod \<mu> \<star> src e\<^sub>0\<guillemotright>"
using assms by auto
show "\<guillemotleft>\<eta>\<^sub>1 \<star> cod \<mu> \<star> \<eta>\<^sub>0 : src e\<^sub>1 \<star> cod \<mu> \<star> src e\<^sub>0 \<Rightarrow> (d\<^sub>1 \<star> e\<^sub>1) \<star> cod \<mu> \<star> (d\<^sub>0 \<star> e\<^sub>0)\<guillemotright>"
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 "\<guillemotleft>(d\<^sub>1 \<star> e\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[cod \<mu>, d\<^sub>0, e\<^sub>0] :
(d\<^sub>1 \<star> e\<^sub>1) \<star> cod \<mu> \<star> d\<^sub>0 \<star> e\<^sub>0 \<Rightarrow> (d\<^sub>1 \<star> e\<^sub>1) \<star> (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
using assms assoc'_in_hom e\<^sub>0.antipar(1-2) e\<^sub>1.antipar(2)
apply (intro hcomp_in_vhom) by auto
show "\<guillemotleft>\<a>[d\<^sub>1, e\<^sub>1, (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0] :
(d\<^sub>1 \<star> e\<^sub>1) \<star> (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0 \<Rightarrow> d\<^sub>1 \<star> e\<^sub>1 \<star> (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
using assms assoc_in_hom e\<^sub>0.antipar(1-2) e\<^sub>1.antipar(2) by auto
show "\<guillemotleft>d\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[e\<^sub>1, cod \<mu> \<star> d\<^sub>0, e\<^sub>0] :
d\<^sub>1 \<star> e\<^sub>1 \<star> (cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0 \<Rightarrow> d\<^sub>1 \<star> (e\<^sub>1 \<star> cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0\<guillemotright>"
using assms assoc'_in_hom [of "d\<^sub>1" "e\<^sub>1 \<star> cod \<mu> \<star> 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 \<phi>_simps [simp]:
assumes "\<guillemotleft>\<mu> : src e\<^sub>0 \<rightarrow> src e\<^sub>1\<guillemotright>"
shows "arr (\<phi> \<mu>)"
and "src (\<phi> \<mu>) = src e\<^sub>0" and "trg (\<phi> \<mu>) = src e\<^sub>1"
and "dom (\<phi> \<mu>) = dom \<mu>" and "cod (\<phi> \<mu>) = d\<^sub>1 \<star> (e\<^sub>1 \<star> cod \<mu> \<star> d\<^sub>0) \<star> e\<^sub>0"
using assms \<phi>_in_hom by auto
interpretation d\<^sub>0: equivalence_in_bicategory V H \<a> \<i> src trg d\<^sub>0 e\<^sub>0 \<open>inv \<epsilon>\<^sub>0\<close> \<open>inv \<eta>\<^sub>0\<close>
using e\<^sub>0.dual_equivalence by simp
interpretation d\<^sub>1: equivalence_in_bicategory V H \<a> \<i> src trg d\<^sub>1 e\<^sub>1 \<open>inv \<epsilon>\<^sub>1\<close> \<open>inv \<eta>\<^sub>1\<close>
using e\<^sub>1.dual_equivalence by simp
interpretation d\<^sub>0e\<^sub>0: two_equivalences_in_bicategory V H \<a> \<i> src trg
d\<^sub>0 e\<^sub>0 \<open>inv \<epsilon>\<^sub>0\<close> \<open>inv \<eta>\<^sub>0\<close> d\<^sub>1 e\<^sub>1 \<open>inv \<epsilon>\<^sub>1\<close> \<open>inv \<eta>\<^sub>1\<close>
..
interpretation \<psi>: inverse_transformation hom'.comp hom'.comp hom'.map \<open>F o G\<close> d\<^sub>0e\<^sub>0.\<phi>
proof -
interpret \<psi>': natural_isomorphism hom'.comp hom'.comp hom'.map \<open>F o G\<close> d\<^sub>0e\<^sub>0.\<phi>
using d\<^sub>0e\<^sub>0.natural_isomorphism_\<phi> e\<^sub>0.antipar e\<^sub>1.antipar d\<^sub>0e\<^sub>0.\<phi>_eq d\<^sub>0e\<^sub>0.\<phi>_def by metis
show "inverse_transformation hom'.comp hom'.comp hom'.map (F o G) d\<^sub>0e\<^sub>0.\<phi>"
..
qed
definition \<psi>
where "\<psi> \<equiv> \<psi>.map"
lemma \<psi>_ide_simp:
assumes "\<guillemotleft>f': trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>" and "ide f'"
shows "\<psi> f' = \<r>[f'] \<cdot> \<l>[f' \<star> trg e\<^sub>0] \<cdot> (\<epsilon>\<^sub>1 \<star> f' \<star> \<epsilon>\<^sub>0) \<cdot> ((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>[f', e\<^sub>0, d\<^sub>0]) \<cdot>
\<a>\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (f' \<star> e\<^sub>0) \<star> d\<^sub>0] \<cdot> (e\<^sub>1 \<star> \<a>[d\<^sub>1, f' \<star> e\<^sub>0, d\<^sub>0])"
proof -
have "hom'.ide f'"
- using assms hom'.ide_char hom'.arr_char by simp
+ using assms hom'.ide_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
hence "\<psi>.map f' = hom'.inv (d\<^sub>0e\<^sub>0.\<phi> f')"
using assms by simp
also have "... = inv (d\<^sub>0e\<^sub>0.\<phi> f')"
- using hom'.inv_char \<open>hom'.ide f'\<close> by simp
+ using hom'.inv_char\<^sub>S\<^sub>b\<^sub>C \<open>hom'.ide f'\<close> by simp
also have "... = inv (d\<^sub>0e\<^sub>0.\<phi>\<^sub>0 f')"
using assms e\<^sub>0.antipar e\<^sub>1.antipar d\<^sub>0e\<^sub>0.\<phi>_eq d\<^sub>0e\<^sub>0.\<phi>_ide_simp [of f'] by simp
also have "... = ((((inv \<r>\<^sup>-\<^sup>1[f'] \<cdot> inv \<l>\<^sup>-\<^sup>1[f' \<star> trg e\<^sub>0]) \<cdot> inv (inv \<epsilon>\<^sub>1 \<star> f' \<star> inv \<epsilon>\<^sub>0)) \<cdot>
inv ((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0])) \<cdot> inv \<a>[e\<^sub>1, d\<^sub>1, (f' \<star> e\<^sub>0) \<star> d\<^sub>0]) \<cdot>
inv (e\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[d\<^sub>1, f' \<star> e\<^sub>0, d\<^sub>0])"
proof -
text \<open>There has to be a better way to do this.\<close>
have 1: "\<And>A B C D E F.
\<lbrakk> iso A; iso B; iso C; iso D; iso E; iso F;
arr (((((A \<cdot> B) \<cdot> C) \<cdot> D) \<cdot> E) \<cdot> F) \<rbrakk> \<Longrightarrow>
inv (((((A \<cdot> B) \<cdot> C) \<cdot> D) \<cdot> E) \<cdot> F) =
inv F \<cdot> inv E \<cdot> inv D \<cdot> inv C \<cdot> inv B \<cdot> inv A"
using inv_comp isos_compose seqE by metis
have "arr (d\<^sub>0e\<^sub>0.\<phi>\<^sub>0 f')"
using assms e\<^sub>0.antipar(2) e\<^sub>1.antipar(2) d\<^sub>0e\<^sub>0.iso_\<phi>\<^sub>0 [of f'] iso_is_arr by simp
moreover have "iso \<r>\<^sup>-\<^sup>1[f']"
using assms iso_runit' by simp
moreover have "iso \<l>\<^sup>-\<^sup>1[f' \<star> trg e\<^sub>0]"
using assms iso_lunit' by auto
moreover have "iso (inv \<epsilon>\<^sub>1 \<star> f' \<star> inv \<epsilon>\<^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 \<star> d\<^sub>1) \<star> \<a>\<^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 \<a>[e\<^sub>1, d\<^sub>1, (f' \<star> e\<^sub>0) \<star> d\<^sub>0]"
using assms e\<^sub>0.antipar e\<^sub>1.antipar by auto
moreover have "iso (e\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[d\<^sub>1, f' \<star> 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 \<star> \<a>\<^sup>-\<^sup>1[d\<^sub>1, f' \<star> e\<^sub>0, d\<^sub>0]" "\<a>[e\<^sub>1, d\<^sub>1, (f' \<star> e\<^sub>0) \<star> d\<^sub>0]"
"(e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0]" "inv \<epsilon>\<^sub>1 \<star> f' \<star> inv \<epsilon>\<^sub>0" "\<l>\<^sup>-\<^sup>1[f' \<star> trg e\<^sub>0]" "\<r>\<^sup>-\<^sup>1[f']"]
comp_assoc
by (metis e\<^sub>0.antipar(2))
qed
also have "... = inv \<r>\<^sup>-\<^sup>1[f'] \<cdot> inv \<l>\<^sup>-\<^sup>1[f' \<star> trg e\<^sub>0] \<cdot> inv (inv \<epsilon>\<^sub>1 \<star> f' \<star> inv \<epsilon>\<^sub>0) \<cdot>
inv ((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0]) \<cdot> inv \<a>[e\<^sub>1, d\<^sub>1, (f' \<star> e\<^sub>0) \<star> d\<^sub>0] \<cdot>
inv (e\<^sub>1 \<star> \<a>\<^sup>-\<^sup>1[d\<^sub>1, f' \<star> e\<^sub>0, d\<^sub>0])"
using comp_assoc by simp
also have "... = \<r>[f'] \<cdot> \<l>[f' \<star> trg e\<^sub>0] \<cdot> (\<epsilon>\<^sub>1 \<star> f' \<star> \<epsilon>\<^sub>0) \<cdot> ((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>[f', e\<^sub>0, d\<^sub>0]) \<cdot>
\<a>\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (f' \<star> e\<^sub>0) \<star> d\<^sub>0] \<cdot> (e\<^sub>1 \<star> \<a>[d\<^sub>1, f' \<star> e\<^sub>0, d\<^sub>0])"
proof -
have "inv \<r>\<^sup>-\<^sup>1[f'] = \<r>[f']"
using assms inv_inv iso_runit by blast
moreover have "inv \<l>\<^sup>-\<^sup>1[f' \<star> trg e\<^sub>0] = \<l>[f' \<star> trg e\<^sub>0]"
using assms iso_lunit by auto
moreover have "inv (inv \<epsilon>\<^sub>1 \<star> f' \<star> inv \<epsilon>\<^sub>0) = \<epsilon>\<^sub>1 \<star> f' \<star> \<epsilon>\<^sub>0"
proof -
have "src (inv \<epsilon>\<^sub>1) = trg f'"
using assms(1) e\<^sub>1.antipar(2) by auto
moreover have "src f' = trg (inv \<epsilon>\<^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 \<psi>_def by simp
qed
lemma \<psi>_components_are_iso:
assumes "\<guillemotleft>f' : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>" and "ide f'"
shows "iso (\<psi> f')"
- using assms \<psi>_def \<psi>.components_are_iso hom'.ide_char hom'.arr_char hom'.iso_char
+ using assms \<psi>_def \<psi>.components_are_iso hom'.ide_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.iso_char\<^sub>S\<^sub>b\<^sub>C
by simp
lemma \<psi>_eq:
shows "\<psi> = (\<lambda>\<mu>'. if \<guillemotleft>\<mu>': trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright> then
\<mu>' \<cdot> \<r>[dom \<mu>'] \<cdot> \<l>[dom \<mu>' \<star> trg e\<^sub>0] \<cdot> (\<epsilon>\<^sub>1 \<star> dom \<mu>' \<star> \<epsilon>\<^sub>0) \<cdot>
((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>[dom \<mu>', e\<^sub>0, d\<^sub>0]) \<cdot> \<a>\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0] \<cdot>
(e\<^sub>1 \<star> \<a>[d\<^sub>1, dom \<mu>' \<star> e\<^sub>0, d\<^sub>0])
else null)"
proof
fix \<mu>'
have "\<not> \<guillemotleft>\<mu>': trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright> \<Longrightarrow> \<psi>.map \<mu>' = null"
- using \<psi>.is_extensional hom'.arr_char hom'.null_char by simp
+ using \<psi>.is_extensional hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.null_char by simp
moreover have "\<guillemotleft>\<mu>': trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright> \<Longrightarrow>
\<psi>.map \<mu>' = \<mu>' \<cdot> \<r>[dom \<mu>'] \<cdot> \<l>[dom \<mu>' \<star> trg e\<^sub>0] \<cdot> (\<epsilon>\<^sub>1 \<star> dom \<mu>' \<star> \<epsilon>\<^sub>0) \<cdot>
((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>[dom \<mu>', e\<^sub>0, d\<^sub>0]) \<cdot> \<a>\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0] \<cdot>
(e\<^sub>1 \<star> \<a>[d\<^sub>1, dom \<mu>' \<star> e\<^sub>0, d\<^sub>0])"
proof -
assume \<mu>': "\<guillemotleft>\<mu>': trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>"
have "\<guillemotleft>\<psi>.map (dom \<mu>') : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>"
- using \<mu>' hom'.arr_char hom'.dom_closed by auto
+ using \<mu>' hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.dom_closed by auto
moreover have "seq \<mu>' (\<psi>.map (dom \<mu>'))"
proof -
have "hom'.seq \<mu>' (\<psi>.map (dom \<mu>'))"
- using \<mu>' \<psi>.preserves_cod hom'.arrI hom'.dom_simp hom'.cod_simp
+ using \<mu>' \<psi>.preserves_cod hom'.arrI\<^sub>S\<^sub>b\<^sub>C hom'.dom_simp hom'.cod_simp
apply (intro hom'.seqI) by auto
thus ?thesis
- using hom'.seq_char by blast
+ using hom'.seq_char\<^sub>S\<^sub>b\<^sub>C by blast
qed
ultimately show ?thesis
- using \<mu>' \<psi>.is_natural_1 [of \<mu>'] hom'.comp_char hom'.arr_char hom'.dom_closed
+ using \<mu>' \<psi>.is_natural_1 [of \<mu>'] hom'.comp_char hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.dom_closed
\<psi>_ide_simp [of "dom \<mu>'"] hom'.dom_simp hom'.cod_simp
apply auto
by (metis \<psi>_def hom'.inclusion ide_dom)
qed
ultimately show "\<psi> \<mu>' = (if \<guillemotleft>\<mu>' : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright> then
\<mu>' \<cdot> \<r>[dom \<mu>'] \<cdot> \<l>[dom \<mu>' \<star> trg e\<^sub>0] \<cdot> (\<epsilon>\<^sub>1 \<star> dom \<mu>' \<star> \<epsilon>\<^sub>0) \<cdot>
((e\<^sub>1 \<star> d\<^sub>1) \<star> \<a>[dom \<mu>', e\<^sub>0, d\<^sub>0]) \<cdot>
\<a>\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0] \<cdot>
(e\<^sub>1 \<star> \<a>[d\<^sub>1, dom \<mu>' \<star> e\<^sub>0, d\<^sub>0])
else null)"
unfolding \<psi>_def by argo
qed
lemma \<psi>_in_hom [intro]:
assumes "\<guillemotleft>\<mu>' : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>"
shows "\<guillemotleft>\<psi> \<mu>' : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>"
and "\<guillemotleft>\<psi> \<mu>' : e\<^sub>1 \<star> (d\<^sub>1 \<star> dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0 \<Rightarrow> cod \<mu>'\<guillemotright>"
proof -
have 0: "\<psi> \<mu>' = \<psi>.map \<mu>'"
using \<psi>_def by auto
hence 1: "hom'.arr (\<psi> \<mu>')"
- using assms hom'.arr_char \<psi>.preserves_reflects_arr by auto
+ using assms hom'.arr_char\<^sub>S\<^sub>b\<^sub>C \<psi>.preserves_reflects_arr by auto
show "\<guillemotleft>\<psi> \<mu>' : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>"
- using 1 hom'.arr_char by blast
+ using 1 hom'.arr_char\<^sub>S\<^sub>b\<^sub>C by blast
show "\<guillemotleft>\<psi> \<mu>' : e\<^sub>1 \<star> (d\<^sub>1 \<star> dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0 \<Rightarrow> cod \<mu>'\<guillemotright>"
- using assms 0 1 \<psi>.preserves_hom hom'.in_hom_char hom'.arr_char
- e\<^sub>0.antipar e\<^sub>1.antipar \<psi>.preserves_dom \<psi>.preserves_cod hom'.dom_char
+ using assms 0 1 \<psi>.preserves_hom hom'.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C
+ e\<^sub>0.antipar e\<^sub>1.antipar \<psi>.preserves_dom \<psi>.preserves_cod hom'.dom_char\<^sub>S\<^sub>b\<^sub>C
apply (intro in_homI)
apply auto[1]
proof -
show "dom (\<psi> \<mu>') = e\<^sub>1 \<star> (d\<^sub>1 \<star> dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0"
proof -
have "hom'.dom (\<psi> \<mu>') = FG.map (hom'.dom \<mu>')"
- using assms 0 \<psi>.preserves_dom hom'.arr_char
+ using assms 0 \<psi>.preserves_dom hom'.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting))
thus ?thesis
- using assms 0 1 hom.arr_char hom'.arr_char hom'.dom_char G.preserves_arr
+ using assms 0 1 hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.dom_char\<^sub>S\<^sub>b\<^sub>C G.preserves_arr
hom'.dom_closed
by auto
qed
show "cod (\<psi> \<mu>') = cod \<mu>'"
proof -
have "hom'.cod (\<psi> \<mu>') = cod \<mu>'"
- using assms 0 \<psi>.preserves_cod hom'.arr_char hom'.cod_simp by auto
+ using assms 0 \<psi>.preserves_cod hom'.arr_char\<^sub>S\<^sub>b\<^sub>C 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
+ using assms 0 1 hom'.arr_char\<^sub>S\<^sub>b\<^sub>C hom'.cod_char\<^sub>S\<^sub>b\<^sub>C G.preserves_arr hom'.cod_closed by auto
qed
qed
qed
lemma \<psi>_simps [simp]:
assumes "\<guillemotleft>\<mu>' : trg e\<^sub>0 \<rightarrow> trg e\<^sub>1\<guillemotright>"
shows "arr (\<psi> \<mu>')"
and "src (\<psi> \<mu>') = trg e\<^sub>0" and "trg (\<psi> \<mu>') = trg e\<^sub>1"
and "dom (\<psi> \<mu>') = e\<^sub>1 \<star> (d\<^sub>1 \<star> dom \<mu>' \<star> e\<^sub>0) \<star> d\<^sub>0" and "cod (\<psi> \<mu>') = cod \<mu>'"
using assms \<psi>_in_hom by auto
interpretation equivalence_of_categories hom'.comp hom.comp F G \<phi> \<psi>
proof -
interpret \<phi>: natural_isomorphism hom.comp hom.comp hom.map \<open>G o F\<close> \<phi>
using \<phi>.natural_isomorphism_axioms \<phi>_def by simp
interpret \<psi>: natural_isomorphism hom'.comp hom'.comp \<open>F o G\<close> hom'.map \<psi>
using \<psi>.natural_isomorphism_axioms \<psi>_def by simp
show "equivalence_of_categories hom'.comp hom.comp F G \<phi> \<psi>"
..
qed
lemma induces_equivalence_of_hom_categories:
shows "equivalence_of_categories hom'.comp hom.comp F G \<phi> \<psi>"
..
lemma equivalence_functor_F:
shows "equivalence_functor hom.comp hom'.comp F"
proof -
interpret \<phi>': inverse_transformation hom.comp hom.comp hom.map \<open>G o F\<close> \<phi> ..
interpret \<psi>': inverse_transformation hom'.comp hom'.comp \<open>F o G\<close> hom'.map \<psi> ..
interpret E: equivalence_of_categories hom.comp hom'.comp G F \<psi>'.map \<phi>'.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 \<open>
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.
\<close>
lemma equivalence_cancel_left:
assumes "equivalence_map e"
and "par \<mu> \<mu>'" and "src e = trg \<mu>" and "e \<star> \<mu> = e \<star> \<mu>'"
shows "\<mu> = \<mu>'"
proof -
obtain d \<eta> \<epsilon> where d\<eta>\<epsilon>: "equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>"
using assms equivalence_map_def by auto
interpret e: equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>
using d\<eta>\<epsilon> by simp
interpret i: equivalence_in_bicategory V H \<a> \<i> src trg
\<open>src \<mu>\<close> \<open>src \<mu>\<close> \<open>inv \<i>[src \<mu>]\<close> \<open>\<i>[src \<mu>]\<close>
using assms iso_unit obj_src
by unfold_locales simp_all
interpret two_equivalences_in_bicategory V H \<a> \<i> src trg
\<open>src \<mu>\<close> \<open>src \<mu>\<close> \<open>inv \<i>[src \<mu>]\<close> \<open>\<i>[src \<mu>]\<close> e d \<eta> \<epsilon>
..
interpret hom: subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (src (src \<mu>)) (src e)\<close>
using hhom_is_subcategory by blast
interpret hom': subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (trg (src \<mu>)) (trg e)\<close>
using hhom_is_subcategory by blast
interpret F: equivalence_functor hom.comp hom'.comp \<open>\<lambda>\<mu>'. e \<star> \<mu>' \<star> src \<mu>\<close>
using equivalence_functor_F by simp
have "F \<mu> = F \<mu>'"
- using assms hom.arr_char
+ using assms hom.arr_char\<^sub>S\<^sub>b\<^sub>C
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 \<mu> \<mu>'"
- using assms hom.arr_char hom.dom_char hom.cod_char
+ using assms hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.dom_char\<^sub>S\<^sub>b\<^sub>C hom.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) in_hhomI src_dom src_src trg_dom)
ultimately show "\<mu> = \<mu>'"
using F.is_faithful by blast
qed
lemma equivalence_cancel_right:
assumes "equivalence_map e"
and "par \<mu> \<mu>'" and "src \<mu> = trg e" and "\<mu> \<star> e = \<mu>' \<star> e"
shows "\<mu> = \<mu>'"
proof -
obtain d \<eta> \<epsilon> where d\<eta>\<epsilon>: "equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>"
using assms equivalence_map_def by auto
interpret e: equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>
using d\<eta>\<epsilon> by simp
interpret i: equivalence_in_bicategory V H \<a> \<i> src trg
\<open>trg \<mu>\<close> \<open>trg \<mu>\<close> \<open>inv \<i>[trg \<mu>]\<close> \<open>\<i>[trg \<mu>]\<close>
using assms iso_unit obj_src
by unfold_locales simp_all
interpret two_equivalences_in_bicategory V H \<a> \<i> src trg
e d \<eta> \<epsilon> \<open>trg \<mu>\<close> \<open>trg \<mu>\<close> \<open>inv \<i>[trg \<mu>]\<close> \<open>\<i>[trg \<mu>]\<close>
..
interpret hom: subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (trg e) (trg (trg \<mu>))\<close>
using hhom_is_subcategory by blast
interpret hom': subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (src e) (src (trg \<mu>))\<close>
using hhom_is_subcategory by blast
interpret G: equivalence_functor hom.comp hom'.comp \<open>\<lambda>\<mu>'. trg \<mu> \<star> \<mu>' \<star> e\<close>
using equivalence_functor_G by simp
have "G \<mu> = G \<mu>'"
- using assms hom.arr_char by simp
+ using assms hom.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "hom.par \<mu> \<mu>'"
- using assms hom.arr_char hom.dom_char hom.cod_char
+ using assms hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.dom_char\<^sub>S\<^sub>b\<^sub>C hom.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) in_hhomI src_dom trg_dom trg_trg)
ultimately show "\<mu> = \<mu>'"
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 \<star> f \<cong> e \<star> f'"
shows "f \<cong> 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 \<eta> \<epsilon> where e: "equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>"
using assms equivalence_map_def by auto
interpret e: equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>
using e by simp
interpret i: equivalence_in_bicategory V H \<a> \<i> src trg
\<open>src f\<close> \<open>src f\<close> \<open>inv \<i>[src f]\<close> \<open>\<i>[src f]\<close>
using assms iso_unit obj_src
by unfold_locales auto
interpret two_equivalences_in_bicategory V H \<a> \<i> src trg
\<open>src f\<close> \<open>src f\<close> \<open>inv \<i>[src f]\<close> \<open>\<i>[src f]\<close> e d \<eta> \<epsilon>
..
interpret hom: subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (src (src f)) (src e)\<close>
using hhom_is_subcategory by blast
interpret hom': subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (trg (src f)) (trg e)\<close>
using hhom_is_subcategory by blast
interpret F: equivalence_functor hom.comp hom'.comp \<open>\<lambda>\<mu>'. e \<star> \<mu>' \<star> src f\<close>
using equivalence_functor_F by simp
have 1: "F f \<cong> F f'"
proof -
have "F f \<cong> (e \<star> f) \<star> src f"
using assms hcomp_assoc_isomorphic equivalence_map_is_ide isomorphic_symmetric
by auto
also have "... \<cong> (e \<star> f') \<star> src f'"
using assms ef' by (simp add: hcomp_isomorphic_ide)
also have "... \<cong> F f'"
using assms ef' hcomp_assoc_isomorphic equivalence_map_is_ide by auto
finally show ?thesis by blast
qed
show "f \<cong> f'"
proof -
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : F f \<Rightarrow> F f'\<guillemotright> \<and> iso \<psi>"
using 1 isomorphic_def by auto
have 2: "hom'.iso \<psi>"
- using assms \<psi> hom'.iso_char hom'.arr_char vconn_implies_hpar(1,2)
+ using assms \<psi> hom'.iso_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C vconn_implies_hpar(1,2)
by auto
have 3: "hom'.in_hom \<psi> (F f) (F f')"
- using assms 2 \<psi> ef' hom'.in_hom_char hom'.arr_char
- by (metis F.preserves_reflects_arr hom'.iso_is_arr hom.arr_char i.antipar(1)
+ using assms 2 \<psi> ef' hom'.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis F.preserves_reflects_arr hom'.iso_is_arr hom.arr_char\<^sub>S\<^sub>b\<^sub>C i.antipar(1)
ideD(1) ide_in_hom(1) trg_src)
obtain \<phi> where \<phi>: "hom.in_hom \<phi> f f' \<and> F \<phi> = \<psi>"
- using assms 3 \<psi> F.is_full F.preserves_reflects_arr hom'.in_hom_char hom.ide_char
+ using assms 3 \<psi> F.is_full F.preserves_reflects_arr hom'.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom.ide_char\<^sub>S\<^sub>b\<^sub>C
by fastforce
have "hom.iso \<phi>"
using 2 \<phi> F.reflects_iso by auto
thus ?thesis
- using \<phi> isomorphic_def hom.in_hom_char hom.arr_char hom.iso_char by auto
+ using \<phi> isomorphic_def hom.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.iso_char\<^sub>S\<^sub>b\<^sub>C 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 \<star> e \<cong> f' \<star> e"
shows "f \<cong> 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 \<eta> \<epsilon> where d\<eta>\<epsilon>: "equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>"
using assms equivalence_map_def by auto
interpret e: equivalence_in_bicategory V H \<a> \<i> src trg e d \<eta> \<epsilon>
using d\<eta>\<epsilon> by simp
interpret i: equivalence_in_bicategory V H \<a> \<i> src trg
\<open>trg f\<close> \<open>trg f\<close> \<open>inv \<i>[trg f]\<close> \<open>\<i>[trg f]\<close>
using assms iso_unit obj_src
by unfold_locales auto
interpret two_equivalences_in_bicategory V H \<a> \<i> src trg
e d \<eta> \<epsilon> \<open>trg f\<close> \<open>trg f\<close> \<open>inv \<i>[trg f]\<close> \<open>\<i>[trg f]\<close>
..
interpret hom: subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (trg e) (trg (trg f))\<close>
using hhom_is_subcategory by blast
interpret hom': subcategory V \<open>\<lambda>\<mu>'. in_hhom \<mu>' (src e) (src (trg f))\<close>
using hhom_is_subcategory by blast
interpret G: equivalence_functor hom.comp hom'.comp \<open>\<lambda>\<mu>'. trg f \<star> \<mu>' \<star> e\<close>
using equivalence_functor_G by simp
have 1: "G f \<cong> G f'"
using assms hcomp_isomorphic_ide hcomp_ide_isomorphic by simp
show "f \<cong> f'"
proof -
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : G f \<Rightarrow> G f'\<guillemotright> \<and> iso \<psi>"
using 1 isomorphic_def by auto
have 2: "hom'.iso \<psi>"
- using assms \<psi> hom'.iso_char hom'.arr_char vconn_implies_hpar(1-2) by auto
+ using assms \<psi> hom'.iso_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C vconn_implies_hpar(1-2) by auto
have 3: "hom'.in_hom \<psi> (G f) (G f')"
- using assms 2 \<psi> 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)
+ using assms 2 \<psi> f'e hom'.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom'.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis G.preserves_arr hom'.iso_is_arr hom.ideI\<^sub>S\<^sub>b\<^sub>C hom.ide_char ideD(1)
ide_in_hom(1) trg_trg)
obtain \<phi> where \<phi>: "hom.in_hom \<phi> f f' \<and> G \<phi> = \<psi>"
- using assms 3 \<psi> G.is_full G.preserves_reflects_arr hom'.in_hom_char hom.ide_char
+ using assms 3 \<psi> G.is_full G.preserves_reflects_arr hom'.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom.ide_char\<^sub>S\<^sub>b\<^sub>C
by fastforce
have "hom.iso \<phi>"
using 2 \<phi> G.reflects_iso by auto
thus ?thesis
- using \<phi> isomorphic_def hom.in_hom_char hom.arr_char hom.iso_char by auto
+ using \<phi> isomorphic_def hom.in_hom_char\<^sub>S\<^sub>b\<^sub>C hom.arr_char\<^sub>S\<^sub>b\<^sub>C hom.iso_char\<^sub>S\<^sub>b\<^sub>C 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 <stark@cs.stonybrook.edu>, 2020
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Modifications"
text \<open>
Modifications are morphisms of pseudonatural transformations.
In this section, we give a definition of ``modification'', and we prove that
the mappings \<open>\<eta>\<close> and \<open>\<epsilon>\<close>, which were chosen in the course of showing that a
pseudonatural equivalence \<open>\<tau>\<close> has a converse pseudonatural equivalence \<open>\<tau>'\<close>,
are invertible modifications that relate the composites \<open>\<tau>'\<tau>\<close> and \<open>\<tau>\<tau>'\<close>
to the identity pseudonatural transformations on \<open>F\<close> and \<open>G\<close>.
This means that \<open>\<tau>\<close> and \<open>\<tau>'\<close> 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.
\<close>
theory Modification
imports PseudonaturalTransformation
begin
locale modification = (* 12 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
\<tau>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1 +
\<tau>': pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0' \<tau>\<^sub>1'
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and \<tau>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>0' :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1' :: "'c \<Rightarrow> 'd"
and \<Gamma> :: "'c \<Rightarrow> 'd" +
assumes \<Gamma>_in_hom: "C.obj a \<Longrightarrow> \<guillemotleft>\<Gamma> a : \<tau>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0' a\<guillemotright>"
and naturality: "\<lbrakk> \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>; C.ide f \<rbrakk> \<Longrightarrow> \<tau>\<^sub>1' f \<cdot>\<^sub>D (G f \<star>\<^sub>D \<Gamma> a) = (\<Gamma> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<tau>\<^sub>1 f"
locale invertible_modification = (* 13 sec *)
modification +
assumes components_are_iso: "C.obj a \<Longrightarrow> D.iso (\<Gamma> a)"
locale identity_modification = (* 12 sec *)
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
\<tau>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and \<tau>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1 :: "'c \<Rightarrow> 'd"
begin
abbreviation map
where "map \<equiv> \<tau>\<^sub>0"
sublocale invertible_modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1 \<tau>\<^sub>0 \<tau>\<^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 \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
\<rho>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<rho>\<^sub>0 \<rho>\<^sub>1 +
\<sigma>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<sigma>\<^sub>0 \<sigma>\<^sub>1 +
\<tau>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1 +
\<Gamma>: modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<rho>\<^sub>0 \<rho>\<^sub>1 \<sigma>\<^sub>0 \<sigma>\<^sub>1 \<Gamma> +
\<Delta>: modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<sigma>\<^sub>0 \<sigma>\<^sub>1 \<tau>\<^sub>0 \<tau>\<^sub>1 \<Delta>
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and \<rho>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<rho>\<^sub>1 :: "'c \<Rightarrow> 'd"
and \<sigma>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<sigma>\<^sub>1 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1 :: "'c \<Rightarrow> 'd"
and \<Gamma> :: "'c \<Rightarrow> 'd"
and \<Delta> :: "'c \<Rightarrow> 'd"
begin
abbreviation map
where "map a \<equiv> \<Delta> a \<cdot>\<^sub>D \<Gamma> a"
sublocale modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<rho>\<^sub>0 \<rho>\<^sub>1 \<tau>\<^sub>0 \<tau>\<^sub>1 map
using \<Delta>.\<Gamma>_in_hom \<Gamma>.\<Gamma>_in_hom \<Delta>.naturality \<Gamma>.naturality
apply unfold_locales
apply auto[1]
proof -
fix f a b
assume f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>" and ide_f: "C.ide f"
have "\<tau>\<^sub>1 f \<cdot>\<^sub>D (G f \<star>\<^sub>D map a) = (\<tau>\<^sub>1 f \<cdot>\<^sub>D (G f \<star>\<^sub>D \<Delta> a)) \<cdot>\<^sub>D (G f \<star>\<^sub>D \<Gamma> a)"
proof -
have "D.seq (\<Delta> a) (\<Gamma> a)"
using f \<Delta>.\<Gamma>_in_hom \<Gamma>.\<Gamma>_in_hom by blast
thus ?thesis
using f ide_f D.whisker_left [of "G f" "\<Delta> a" "\<Gamma> a"] D.comp_assoc
by simp
qed
also have "... = (\<Delta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<sigma>\<^sub>1 f \<cdot>\<^sub>D (G f \<star>\<^sub>D \<Gamma> a)"
using f ide_f \<Delta>.naturality [of f a b] D.comp_assoc by simp
also have "... = ((\<Delta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Gamma> b \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<rho>\<^sub>1 f"
using f ide_f \<Gamma>.naturality [of f a b] D.comp_assoc by simp
also have "... = (map b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<rho>\<^sub>1 f"
proof -
have "D.seq (\<Delta> b) (\<Gamma> b)"
using f \<Delta>.\<Gamma>_in_hom \<Gamma>.\<Gamma>_in_hom by blast
thus ?thesis
using f ide_f D.whisker_right [of "F f" "\<Delta> b" "\<Gamma> b"] by simp
qed
finally show "\<tau>\<^sub>1 f \<cdot>\<^sub>D (G f \<star>\<^sub>D map a) = (map b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<rho>\<^sub>1 f" by simp
qed
end
context converse_pseudonatural_equivalence
begin
interpretation EV: self_evaluation_map V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
notation EV.eval ("\<lbrace>_\<rbrace>")
interpretation \<iota>\<^sub>F: identity_pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F
..
interpretation \<iota>\<^sub>G: identity_pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
..
interpretation \<tau>'\<tau>: composite_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G F \<Phi>\<^sub>F
\<tau>\<^sub>0 \<tau>\<^sub>1 map\<^sub>0 map\<^sub>1
..
interpretation \<tau>\<tau>': composite_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G F \<Phi>\<^sub>F G \<Phi>\<^sub>G
map\<^sub>0 map\<^sub>1 \<tau>\<^sub>0 \<tau>\<^sub>1
..
interpretation \<eta>: invertible_modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F F \<Phi>\<^sub>F
\<iota>\<^sub>F.map\<^sub>0 \<iota>\<^sub>F.map\<^sub>1 \<tau>'\<tau>.map\<^sub>0 \<tau>'\<tau>.map\<^sub>1 \<eta>
proof
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>\<eta> a : F.map\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>'\<tau>.map\<^sub>0 a\<guillemotright>"
using unit_in_hom \<tau>'\<tau>.map\<^sub>0_def by simp
show "\<And>a. C.obj a \<Longrightarrow> D.iso (\<eta> a)"
by simp
show "\<And>f a b. \<lbrakk> \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>; C.ide f\<rbrakk>
\<Longrightarrow> \<tau>'\<tau>.map\<^sub>1 f \<cdot>\<^sub>D (F f \<star>\<^sub>D \<eta> a) = (\<eta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
proof -
fix f a b
assume f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>" and ide_f: "C.ide f"
have a: "C.obj a" and b: "C.obj b"
using f by auto
have "\<tau>'\<tau>.map\<^sub>1 f \<cdot>\<^sub>D (F f \<star>\<^sub>D \<eta> a) =
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a)
\<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<eta> a)"
unfolding \<tau>'\<tau>.map\<^sub>1_def map\<^sub>1_def
using a b f D.comp_assoc by auto
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<eta> a)"
proof -
have "(\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a)
\<star>\<^sub>D \<tau>\<^sub>0 a
= ((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "D.arr ((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a))"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
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 "\<tau>\<^sub>0 a"] by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<eta> a)"
proof -
have "((\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] =
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a b f ide_f D.assoc'_naturality [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F f]" "\<tau>\<^sub>0' a" "\<tau>\<^sub>0 a"] by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
((F\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<eta> a) = \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<eta> a"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F f]" "F f" "\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a" "\<eta> a"]
by simp
also have "... = ((F\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^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 \<star>\<^sub>D F f" "\<l>\<^sub>D\<^sup>-\<^sup>1[F f]" "\<eta> a" "F\<^sub>0 a"]
by auto
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<eta> a) =
((F\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((F\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] =
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D ((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a b f ide_f D.assoc'_naturality [of "\<eta> b \<star>\<^sub>D F f" "\<tau>\<^sub>0' a" "\<tau>\<^sub>0 a"] by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D ((F\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a)
= (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> b \<star>\<^sub>D F f" "F\<^sub>0 b \<star>\<^sub>D F f" "\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a" "\<eta> a"]
by auto
also have "... = (((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D ((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a)"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f" "\<eta> b \<star>\<^sub>D F f" "\<eta> a" "F\<^sub>0 a"]
by auto
finally have "((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D ((F\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) =
(((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D ((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) =
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a"
using a b f ide_f D.whisker_right \<tau>.iso_map\<^sub>1_ide by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 a), \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a)) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] =
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 a), \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
D.assoc'_naturality
[of "(\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]" "\<tau>\<^sub>0' a" "\<tau>\<^sub>0 a"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 a), \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) =
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]) \<cdot>\<^sub>D ((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<eta> a"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
D.interchange [of "(\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]"
"(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f" "\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a" "\<eta> a"]
by auto
also have "... = (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<eta> a"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.comp_assoc
by auto
also have "... = (\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]
\<star>\<^sub>D \<eta> a \<cdot>\<^sub>D F\<^sub>0 a"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr by auto
also have "... = ((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "D.seq (\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) ((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f])"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
by (intro D.seqI D.hseqI') auto
thus ?thesis
using a b f ide_f \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a)"
"(\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]"
"\<eta> a" "F\<^sub>0 a"]
by simp
qed
finally have "((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f) \<star>\<^sub>D \<eta> a) =
((\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 a), \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a
= (\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a)"
proof -
have "(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a
= (\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a b f ide_f D.hcomp_reassoc D.whisker_right [of "\<tau>\<^sub>0 a"] by auto
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a b f ide_f D.hcomp_reassoc(1) [of "\<tau>\<^sub>0' b \<star>\<^sub>D G f" "\<epsilon> a" "\<tau>\<^sub>0 a"]
by auto
finally show ?thesis by blast
qed
moreover have "(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<eta> a
= (\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<eta> a =
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D ((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<eta> a \<cdot>\<^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 "\<tau>\<^sub>0' b" "G f" "\<tau>\<^sub>0 a"]
by auto
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^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 "... = (\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a)"
using a b f ide_f D.hcomp_reassoc(1) [of "\<tau>\<^sub>0' b \<star>\<^sub>D G f" "\<tau>\<^sub>0 a" "\<eta> a"] by auto
finally show ?thesis by blast
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 a), \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] =
(\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]"
proof -
have "\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 a), \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a]
= \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[ \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>G f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>G f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>), \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>]\<rbrace>"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
also have "... = \<lbrace>(\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>]\<rbrace>"
using a b f ide_f by (intro EV.eval_eqI, auto)
also have "... = (\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using a chosen_adjoint_equivalence by simp
have "((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a)
= (\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D (\<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a)"
using a b f ide_f D.whisker_left by auto
also have "... = (\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]"
using triangle_left by simp
finally have "((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a, \<tau>\<^sub>0' a, \<tau>\<^sub>0 a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<eta> a)
= (\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
(\<r>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a)) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a)
= \<r>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a]"
proof -
have "\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, G f, G\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D G f, G\<^sub>0 a, \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D G f) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f, \<tau>\<^sub>0 a, F\<^sub>0 a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, G f, \<tau>\<^sub>0 a] \<star>\<^sub>D F\<^sub>0 a)
= \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<r>\<^bold>[\<^bold>\<langle>G f\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>G\<^sub>0 a\<^bold>\<rangle>\<^sub>0\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>G\<^sub>0 a\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<r>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 a\<^bold>\<rangle>\<^sub>0\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>G f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>F\<^sub>0 a\<^bold>\<rangle>\<^sub>0)\<rbrace>"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char D.\<ll>_ide_simp D.\<rr>_ide_simp
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.\<ll>_ide_simp D.\<rr>_ide_simp
by auto
also have "... = \<lbrace>\<^bold>\<r>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 a\<^bold>\<rangle>\<^bold>]\<rbrace>"
using a b f ide_f by (intro EV.eval_eqI, auto)
also have "... = \<r>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a]"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char D.\<ll>_ide_simp D.\<rr>_ide_simp
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.\<ll>_ide_simp D.\<rr>_ide_simp
by auto
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)))) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D \<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f] \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "\<r>\<^sub>D[\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D F\<^sub>0 a)
= (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
\<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f]"
using a b f ide_f D.comp_assoc \<tau>.iso_map\<^sub>1_ide
D.runit_naturality [of "(\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]) \<cdot>\<^sub>D
\<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f]) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D ((\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)))
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]"
using a b f ide_f D.comp_arr_inv' D.comp_arr_dom \<tau>.iso_map\<^sub>1_ide
D.whisker_left [of "\<tau>\<^sub>0' b" "\<tau>\<^sub>1 f" "D.inv (\<tau>\<^sub>1 f)"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f] \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f]) \<cdot>\<^sub>D \<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f]
= \<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^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 "... = (\<eta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<r>\<^sub>D[F\<^sub>0 b \<star>\<^sub>D F f] \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D F\<^sub>0 a)"
proof -
have "\<r>\<^sub>D[(\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f] \<cdot>\<^sub>D ((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D F\<^sub>0 a)
= (\<eta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<r>\<^sub>D[F\<^sub>0 b \<star>\<^sub>D F f]"
using a b f ide_f D.runit_naturality [of "\<eta> b \<star>\<^sub>D F f"] by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<eta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
using a b f ide_f D.runit_naturality [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F f]"] by auto
finally show "\<tau>'\<tau>.map\<^sub>1 f \<cdot>\<^sub>D (F f \<star>\<^sub>D \<eta> a) = (\<eta> b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
by blast
qed
qed
lemma unit_is_invertible_modification:
shows "invertible_modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F F \<Phi>\<^sub>F
\<iota>\<^sub>F.map\<^sub>0 \<iota>\<^sub>F.map\<^sub>1 \<tau>'\<tau>.map\<^sub>0 \<tau>'\<tau>.map\<^sub>1 \<eta>"
..
interpretation \<epsilon>: invertible_modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G G \<Phi>\<^sub>G
\<tau>\<tau>'.map\<^sub>0 \<tau>\<tau>'.map\<^sub>1 \<iota>\<^sub>G.map\<^sub>0 \<iota>\<^sub>G.map\<^sub>1 \<epsilon>
proof
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>\<epsilon> a : \<tau>\<tau>'.map\<^sub>0 a \<Rightarrow>\<^sub>D G\<^sub>0 a\<guillemotright>"
using counit_in_hom \<tau>\<tau>'.map\<^sub>0_def by simp
show "\<And>a. C.obj a \<Longrightarrow> D.iso (\<epsilon> a)"
by simp
show "\<And>f a b. \<lbrakk>\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>; C.ide f\<rbrakk>
\<Longrightarrow> (\<l>\<^sub>D\<^sup>-\<^sup>1[G f] \<cdot>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> a) = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D \<tau>\<tau>'.map\<^sub>1 f"
proof -
fix f a b
assume f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>" and ide_f: "C.ide f"
have a: "C.obj a" and b: "C.obj b"
using f by auto
have "(\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D \<tau>\<tau>'.map\<^sub>1 f
= (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
unfolding \<tau>\<tau>'.map\<^sub>1_def map\<^sub>1_def
using a b f D.comp_assoc by auto
also have "... = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "\<tau>\<^sub>0 b \<star>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "D.arr ((\<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a))"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
by (intro D.seqI) auto
thus ?thesis
using a b f D.whisker_left [of "\<tau>\<^sub>0 b"] by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a b f ide_f D.whisker_left [of "\<tau>\<^sub>0 b"] \<tau>.iso_map\<^sub>1_ide by auto
also have "... = \<tau>\<^sub>0 b \<star>\<^sub>D (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
D.assoc_naturality [of "\<tau>\<^sub>0' b" "D.inv (\<tau>\<^sub>1 f)" "\<tau>\<^sub>0' a"]
by auto
also have "... = (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a])"
using a b f ide_f D.whisker_left [of "\<tau>\<^sub>0 b"] \<tau>.iso_map\<^sub>1_ide by auto
finally have "(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a b f ide_f D.whisker_left \<tau>.iso_map\<^sub>1_ide by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f \<star>\<^sub>D G\<^sub>0 a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f] \<cdot>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f \<star>\<^sub>D G\<^sub>0 a]"
using a b f ide_f D.assoc'_naturality [of "\<tau>\<^sub>0 b" "\<tau>\<^sub>0' b" "\<r>\<^sub>D[G f]"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f \<star>\<^sub>D G\<^sub>0 a] \<cdot>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a)
= ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a]"
using a b f ide_f D.assoc'_naturality [of "\<tau>\<^sub>0 b" "\<tau>\<^sub>0' b" "G f \<star>\<^sub>D \<epsilon> a"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a))
= ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a]"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
D.assoc'_naturality
[of "\<tau>\<^sub>0 b" "\<tau>\<^sub>0' b" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "(\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<r>\<^sub>D[G f]) = \<epsilon> b \<star>\<^sub>D \<r>\<^sub>D[G f]"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<epsilon> b" "\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b" "G f" "\<r>\<^sub>D[G f]"]
by simp
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D (\<epsilon> b \<star>\<^sub>D G f \<star>\<^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" "\<epsilon> b" "\<r>\<^sub>D[G f]" "G f \<star>\<^sub>D G\<^sub>0 a"]
by auto
finally have "(\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<r>\<^sub>D[G f]) =
(G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D (\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "(\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 a) \<cdot>\<^sub>D ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a)
= \<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<epsilon> b" "\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b" "G f \<star>\<^sub>D G\<^sub>0 a" "G f \<star>\<^sub>D \<epsilon> a"]
by auto
also have "... = (G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D (\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.interchange [of "G\<^sub>0 b" "\<epsilon> b" "G f \<star>\<^sub>D \<epsilon> a" "G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a"]
by auto
finally have "(\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 a) \<cdot>\<^sub>D ((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) =
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D (\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<epsilon> b \<star>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have 1: "D.seq (G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a)
(\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a))"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
by (intro D.seqI D.hseqI') auto
have 2: "D.seq (\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a))
((\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
by (intro D.seqI D.hseqI') auto
have "(\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) =
\<epsilon> b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using 1 a b f ide_f D.comp_arr_dom D.comp_cod_arr \<tau>.iso_map\<^sub>1_ide
D.interchange [of "\<epsilon> b" "\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b" "G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"]
by auto
also have "... = G\<^sub>0 b \<cdot>\<^sub>D \<epsilon> b \<star>\<^sub>D
(\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr by auto
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<epsilon> b \<star>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using 2 a b f ide_f \<tau>.iso_map\<^sub>1_ide D.comp_assoc
D.interchange [of "G\<^sub>0 b" "\<epsilon> b"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
"(\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a"]
by simp
finally have "(\<epsilon> b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a))
= (G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<epsilon> b \<star>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "\<epsilon> b \<star>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a
= (G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a])"
proof -
have "\<epsilon> b \<star>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a
= (G\<^sub>0 b \<cdot>\<^sub>D \<epsilon> b \<cdot>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b)) \<star>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]"
using a b f ide_f D.comp_arr_dom D.comp_cod_arr
D.hcomp_reassoc(1) [of "\<tau>\<^sub>0 b" "F f" "\<tau>\<^sub>0' a"]
by auto
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^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" "\<epsilon> b \<cdot>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b)" "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]"
"(\<tau>\<^sub>0 b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]"]
D.interchange [of "\<epsilon> b" "\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b" "\<tau>\<^sub>0 b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a"
"\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]"]
by fastforce
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a])"
using a b f ide_f D.hcomp_reassoc(2) [of "\<epsilon> b" "\<tau>\<^sub>0 b" "F f \<star>\<^sub>D \<tau>\<^sub>0' a"]
by auto
finally show ?thesis by blast
qed
moreover have "\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a
= (\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a])"
proof -
have "\<tau>\<^sub>0 b \<star>\<^sub>D (\<eta> b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a =
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a])"
using a b f ide_f D.hcomp_reassoc(1) [of "\<eta> b" "F f" "\<tau>\<^sub>0' a"]
D.whisker_left [of "\<tau>\<^sub>0 b"]
by auto
also have "... = (\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a])"
using a b f ide_f D.hcomp_reassoc(2) [of "\<tau>\<^sub>0 b" "\<eta> b" "F f \<star>\<^sub>D \<tau>\<^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 \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(((\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, \<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, \<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<tau>\<^sub>0' b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, (\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b \<star>\<^sub>D F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' b, \<tau>\<^sub>0 b, F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b \<star>\<^sub>D \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]
= \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]\<rbrace>"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
also have "... = \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<rbrace>"
using a b f ide_f by (intro EV.eval_eqI, auto)
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, \<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by auto
qed
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
((G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "((\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, \<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 b\<close> \<open>\<tau>\<^sub>0' b\<close> \<open>\<eta> b\<close> \<open>\<epsilon> b\<close>
using b chosen_adjoint_equivalence by simp
have "((\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, \<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<epsilon> b \<star>\<^sub>D \<tau>\<^sub>0 b) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, \<tau>\<^sub>0' b, \<tau>\<^sub>0 b] \<cdot>\<^sub>D (\<tau>\<^sub>0 b \<star>\<^sub>D \<eta> b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a"
using a b f ide_f D.whisker_right by auto
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^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 \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]
= \<l>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a]"
proof -
have "(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[G\<^sub>0 b, \<tau>\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 b, F\<^sub>0 b, F f \<star>\<^sub>D \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 b, F f, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0 b \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 b, F f, \<tau>\<^sub>0' a]
= \<lbrace>(\<^bold>\<langle>G\<^sub>0 b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>G\<^sub>0 b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<r>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>F\<^sub>0 b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]\<rbrace>"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char D.\<ll>_ide_simp D.\<rr>_ide_simp
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.\<ll>_ide_simp D.\<rr>_ide_simp
by auto
also have "... = \<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[(\<^bold>\<langle>\<tau>\<^sub>0 b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' a\<^bold>\<rangle>\<^bold>]\<rbrace>"
using a b f ide_f by (intro EV.eval_eqI, auto)
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a]"
- using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char D.\<ll>_ide_simp
+ using a b f ide_f D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.\<ll>_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 \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "(G\<^sub>0 b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0 b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' a]
= \<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide
D.lunit'_naturality [of "\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a)"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D
\<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
proof -
have "((D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (\<tau>\<^sub>1 f \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] =
((D.inv (\<tau>\<^sub>1 f) \<cdot>\<^sub>D \<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
using a b f ide_f \<tau>.iso_map\<^sub>1_ide D.whisker_right [of "\<tau>\<^sub>0' a"] by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
using a b f ide_f \<tau>.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 \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a]"
proof -
have "\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] \<cdot>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] =
\<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^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 \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D G\<^sub>0 a]) \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> a)"
proof -
have "(G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D (G\<^sub>0 b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a] =
(G\<^sub>0 b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G f \<star>\<^sub>D G\<^sub>0 a] \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> a)"
using a b f ide_f D.lunit'_naturality [of "G f \<star>\<^sub>D \<epsilon> a"] by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<l>\<^sub>D\<^sup>-\<^sup>1[G f] \<cdot>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> a)"
using a b f ide_f D.lunit'_naturality [of "\<r>\<^sub>D[G f]"] by auto
finally show "(\<l>\<^sub>D\<^sup>-\<^sup>1[G f] \<cdot>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> a) = (\<epsilon> b \<star>\<^sub>D G f) \<cdot>\<^sub>D \<tau>\<tau>'.map\<^sub>1 f"
by simp
qed
qed
lemma counit_is_invertible_modification:
shows "invertible_modification
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G G \<Phi>\<^sub>G
\<tau>\<tau>'.map\<^sub>0 \<tau>\<tau>'.map\<^sub>1 \<iota>\<^sub>G.map\<^sub>0 \<iota>\<^sub>G.map\<^sub>1 \<epsilon>"
..
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,2959 +1,2959 @@
(* Title: PreBicategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
text \<open>
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 \<open>src\<close> and \<open>trg\<close> that assign to each arrow its ``source''
and ``target'' subject to certain commutativity constraints,
a partial binary operation \<open>\<star>\<close> of horizontal composition that is suitably functorial on
the ``hom-categories'' determined by the assignment of sources and targets,
``associativity'' isomorphisms \<open>\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> (g \<star> h)\<guillemotright>\<close> for each horizontally
composable triple of vertical identities \<open>f\<close>, \<open>g\<close>, \<open>h\<close>, subject to the usual naturality
and coherence conditions, and for each ``object'' \<open>a\<close> (defined to be an arrow that is
its own source and target) a ``unit isomorphism'' \<open>\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>\<close>.
As is the case for monoidal categories, the unit isomorphisms and associator isomorphisms
together enable a canonical definition of left and right ``unit'' isomorphisms
\<open>\<guillemotleft>\<l>[f] : a \<star> f \<Rightarrow> f\<guillemotright>\<close> and \<open>\<guillemotleft>\<r>[f] : f \<star> a \<Rightarrow> f\<guillemotright>\<close> when \<open>f\<close> is a vertical identity
horizontally composable on the left or right by \<open>a\<close>, 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 \<open>\<mu>\<close> and \<open>\<nu>\<close>
are composable if the chosen source of \<open>\<mu>\<close> coincides with the chosen target of \<open>\<nu>\<close>.
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 \<open>\<star>\<close>
on the arrow type of a ``vertical'' category \<open>V\<close>, 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 \<open>V\<close> 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 \<open>a\<close> such that \<open>a \<star> a \<cong> a\<close>
and is such that the mappings \<open>a \<star> \<hyphen>\<close> and \<open>\<hyphen> \<star> a\<close> are fully faithful endofunctors
of the subcategories of \<open>V\<close> consisting of the arrows for which they are defined.
We define the \emph{sources} of an arrow \<open>\<mu>\<close> to be the weak units that are horizontally
composable with \<open>\<mu>\<close> on the right, and the \emph{targets} of \<open>\<mu>\<close> to be the weak units
that are horizontally composable with \<open>\<mu>\<close> on the left.
An \emph{associative weak composition} is defined to be a weak composition that is equipped
with ``associator'' isomorphisms \<open>\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> (g \<star> h)\<guillemotright>\<close> for horizontally
composable vertical identities \<open>f\<close>, \<open>g\<close>, \<open>h\<close>, 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 \<open>\<mu>\<close> and \<open>\<nu>\<close> is characterized by the set of sources of \<open>\<mu>\<close> being
equal to the set of targets of \<open>\<nu>\<close>.
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} \<open>\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>\<close>.
This choice, together with the associator isomorphisms, enables a canonical definition
of left and right unit isomorphisms \<open>\<guillemotleft>\<l>[f] : a \<star> f \<Rightarrow> f\<guillemotright>\<close> and \<open>\<guillemotleft>\<r>[f] : f \<star> a \<Rightarrow> f\<guillemotright>\<close>
when \<open>f\<close> is a vertical identity horizontally composable on the left or right by \<open>a\<close>,
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.
\<close>
theory Prebicategory
imports Category3.EquivalenceOfCategories Category3.Subcategory IsomorphismClass
begin
section "Weak Composition"
text \<open>
In this section we define a locale \<open>weak_composition\<close>, 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
\<open>weak_composition\<close> 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 \<open>\<mu>\<close>,
the sets of arrows horizontally composable on the left and on the right with \<open>\<mu>\<close>
form subcategories with respect to vertical composition. We define the
sources of \<open>\<mu>\<close> to be the weak units that are composable with \<open>\<mu>\<close> on the right,
and the targets of \<open>\<mu>\<close> to be the weak units that are composable with \<open>\<mu>\<close>
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).
\<close>
subsection "Definition"
locale weak_composition =
category V +
VxV: product_category V V +
VoV: subcategory VxV.comp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu> \<noteq> null\<close> +
"functor" VoV.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a comp" (infixr "\<star>" 53) +
assumes left_connected: "seq \<nu> \<nu>' \<Longrightarrow> \<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu>' \<star> \<mu> \<noteq> null"
and right_connected: "seq \<mu> \<mu>' \<Longrightarrow> \<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu> \<star> \<mu>' \<noteq> null"
and match_1: "\<lbrakk> \<nu> \<star> \<mu> \<noteq> null; (\<nu> \<star> \<mu>) \<star> \<tau> \<noteq> null \<rbrakk> \<Longrightarrow> \<mu> \<star> \<tau> \<noteq> null"
and match_2: "\<lbrakk> \<nu> \<star> (\<mu> \<star> \<tau>) \<noteq> null; \<mu> \<star> \<tau> \<noteq> null \<rbrakk> \<Longrightarrow> \<nu> \<star> \<mu> \<noteq> null"
and match_3: "\<lbrakk> \<mu> \<star> \<tau> \<noteq> null; \<nu> \<star> \<mu> \<noteq> null \<rbrakk> \<Longrightarrow> (\<nu> \<star> \<mu>) \<star> \<tau> \<noteq> null"
and match_4: "\<lbrakk> \<mu> \<star> \<tau> \<noteq> null; \<nu> \<star> \<mu> \<noteq> null \<rbrakk> \<Longrightarrow> \<nu> \<star> (\<mu> \<star> \<tau>) \<noteq> null"
begin
text \<open>
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 \<mu>"}
will mean, essentially, ``\<open>\<mu>\<close> 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.
\<close>
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation in_hom ("\<guillemotleft>_ : _ \<Rightarrow> _\<guillemotright>")
lemma is_partial_magma:
shows "partial_magma H"
proof
show "\<exists>!n. \<forall>f. n \<star> f = n \<and> f \<star> n = n"
proof
show 1: "\<forall>f. null \<star> f = null \<and> f \<star> null = null"
- using is_extensional VoV.inclusion VoV.arr_char by force
+ using is_extensional VoV.inclusion VoV.arr_char\<^sub>S\<^sub>b\<^sub>C by force
show "\<And>n. \<forall>f. n \<star> f = n \<and> f \<star> n = n \<Longrightarrow> n = null"
- using 1 VoV.arr_char is_extensional not_arr_null by metis
+ using 1 VoV.arr_char\<^sub>S\<^sub>b\<^sub>C is_extensional not_arr_null by metis
qed
qed
interpretation H: partial_magma H
using is_partial_magma by auto
text \<open>
Either \<open>match_1\<close> or \<open>match_2\<close> seems essential for the next result, which states
that the nulls for the horizontal and vertical compositions coincide.
\<close>
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 "\<nu> \<star> \<mu> \<noteq> null"
shows "arr \<mu>" and "arr \<nu>"
- using assms is_extensional VoV.arr_char VoV.inclusion by auto
+ using assms is_extensional VoV.arr_char\<^sub>S\<^sub>b\<^sub>C VoV.inclusion by auto
lemma hcomp_null [simp]:
shows "null \<star> \<mu> = null" and "\<mu> \<star> null = null"
using H.comp_null by auto
lemma hcomp_simps\<^sub>W\<^sub>C [simp]:
assumes "\<nu> \<star> \<mu> \<noteq> null"
shows "arr (\<nu> \<star> \<mu>)" and "dom (\<nu> \<star> \<mu>) = dom \<nu> \<star> dom \<mu>" and "cod (\<nu> \<star> \<mu>) = cod \<nu> \<star> cod \<mu>"
- using assms preserves_arr preserves_dom preserves_cod VoV.arr_char VoV.inclusion
+ using assms preserves_arr preserves_dom preserves_cod VoV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<star> f \<noteq> null"
shows "ide (g \<star> f)"
- using assms preserves_ide VoV.ide_char by force
+ using assms preserves_ide VoV.ide_char\<^sub>S\<^sub>b\<^sub>C by force
lemma hcomp_in_hom\<^sub>W\<^sub>C [intro]:
assumes "\<nu> \<star> \<mu> \<noteq> null"
shows "\<guillemotleft>\<nu> \<star> \<mu> : dom \<nu> \<star> dom \<mu> \<Rightarrow> cod \<nu> \<star> cod \<mu>\<guillemotright>"
using assms by auto
text \<open>
Horizontal composability of arrows is determined by horizontal composability of
their domains and codomains (defined with respect to vertical composition).
\<close>
lemma hom_connected:
shows "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> dom \<nu> \<star> \<mu> \<noteq> null"
and "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu> \<star> dom \<mu> \<noteq> null"
and "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> cod \<nu> \<star> \<mu> \<noteq> null"
and "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu> \<star> cod \<mu> \<noteq> null"
proof -
show "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> dom \<nu> \<star> \<mu> \<noteq> null"
using left_connected [of \<nu> "dom \<nu>" \<mu>] composable_implies_arr arr_dom_iff_arr by force
show "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> cod \<nu> \<star> \<mu> \<noteq> null"
using left_connected [of "cod \<nu>" \<nu> \<mu>] composable_implies_arr arr_cod_iff_arr by force
show "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu> \<star> dom \<mu> \<noteq> null"
using right_connected [of \<mu> "dom \<mu>" \<nu>] composable_implies_arr arr_dom_iff_arr by force
show "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu> \<star> cod \<mu> \<noteq> null"
using right_connected [of "cod \<mu>" \<mu> \<nu>] composable_implies_arr arr_cod_iff_arr by force
qed
lemma isomorphic_implies_equicomposable:
assumes "f \<cong> g"
shows "\<tau> \<star> f \<noteq> null \<longleftrightarrow> \<tau> \<star> g \<noteq> null"
and "f \<star> \<sigma> \<noteq> null \<longleftrightarrow> g \<star> \<sigma> \<noteq> null"
using assms isomorphic_def hom_connected by auto
lemma interchange:
assumes "seq \<nu> \<mu>" and "seq \<tau> \<sigma>"
shows "(\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>) = (\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)"
proof -
have "\<mu> \<star> \<sigma> = null \<Longrightarrow> ?thesis"
by (metis assms comp_null(2) dom_comp hom_connected(1-2))
moreover have "\<mu> \<star> \<sigma> \<noteq> null \<Longrightarrow> ?thesis"
proof -
assume \<mu>\<sigma>: "\<mu> \<star> \<sigma> \<noteq> null"
have 1: "VoV.arr (\<mu>, \<sigma>)"
- using \<mu>\<sigma> VoV.arr_char by auto
+ using \<mu>\<sigma> VoV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
have \<nu>\<tau>: "(\<nu>, \<tau>) \<in> VoV.hom (VoV.cod (\<mu>, \<sigma>)) (VoV.cod (\<nu>, \<tau>))"
proof -
have "VoV.arr (\<nu>, \<tau>)"
- using assms 1 hom_connected VoV.arr_char
+ using assms 1 hom_connected VoV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (elim seqE, auto, metis)
thus ?thesis
- using assms \<mu>\<sigma> VoV.dom_char VoV.cod_char by fastforce
+ using assms \<mu>\<sigma> VoV.dom_char\<^sub>S\<^sub>b\<^sub>C VoV.cod_char\<^sub>S\<^sub>b\<^sub>C by fastforce
qed
show ?thesis
proof -
have "VoV.seq (\<nu>, \<tau>) (\<mu>, \<sigma>)"
using assms 1 \<mu>\<sigma> \<nu>\<tau> VoV.seqI by blast
thus ?thesis
using assms 1 \<mu>\<sigma> \<nu>\<tau> VoV.comp_char preserves_comp [of "(\<nu>, \<tau>)" "(\<mu>, \<sigma>)"] VoV.seqI
by fastforce
qed
qed
ultimately show ?thesis by blast
qed
lemma paste_1:
shows "\<nu> \<star> \<mu> = (cod \<nu> \<star> \<mu>) \<cdot> (\<nu> \<star> dom \<mu>)"
using interchange composable_implies_arr comp_arr_dom comp_cod_arr
hom_connected(2-3)
by (metis comp_null(2))
lemma paste_2:
shows "\<nu> \<star> \<mu> = (\<nu> \<star> cod \<mu>) \<cdot> (dom \<nu> \<star> \<mu>)"
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 \<nu> \<mu>" and "ide f"
shows "f \<star> (\<nu> \<cdot> \<mu>) = (f \<star> \<nu>) \<cdot> (f \<star> \<mu>)"
using assms interchange [of f f \<nu> \<mu>] hom_connected by auto
lemma whisker_right:
assumes "seq \<nu> \<mu>" and "ide f"
shows "(\<nu> \<cdot> \<mu>) \<star> f = (\<nu> \<star> f) \<cdot> (\<mu> \<star> f)"
using assms interchange [of \<nu> \<mu> f f] hom_connected by auto
subsection "Hom-Subcategories"
definition left
where "left \<tau> \<equiv> \<lambda>\<mu>. \<tau> \<star> \<mu> \<noteq> null"
definition right
where "right \<sigma> \<equiv> \<lambda>\<mu>. \<mu> \<star> \<sigma> \<noteq> null"
lemma right_iff_left:
shows "right \<sigma> \<tau> \<longleftrightarrow> left \<tau> \<sigma>"
using right_def left_def by simp
lemma left_respects_isomorphic:
assumes "f \<cong> g"
shows "left f = left g"
using assms isomorphic_implies_equicomposable left_def by auto
lemma right_respects_isomorphic:
assumes "f \<cong> g"
shows "right f = right g"
using assms isomorphic_implies_equicomposable right_def by auto
lemma left_iff_left_inv:
assumes "iso \<mu>"
shows "left \<tau> \<mu> \<longleftrightarrow> left \<tau> (inv \<mu>)"
using assms left_def inv_in_hom hom_connected(2) hom_connected(4) [of \<tau> "inv \<mu>"]
by auto
lemma right_iff_right_inv:
assumes "iso \<mu>"
shows "right \<sigma> \<mu> \<longleftrightarrow> right \<sigma> (inv \<mu>)"
using assms right_def inv_in_hom hom_connected(1) hom_connected(3) [of "inv \<mu>" \<sigma>]
by auto
lemma left_hom_is_subcategory:
assumes "arr \<mu>"
shows "subcategory V (left \<mu>)"
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 \<mu>"
shows "subcategory V (right \<mu>)"
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 \<equiv> subcategory.comp V (left a)"
abbreviation Right
where "Right a \<equiv> subcategory.comp V (right a)"
text \<open>
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.
\<close>
definition H\<^sub>L
where "H\<^sub>L g \<equiv> \<lambda>\<mu>. g \<star> \<mu>"
definition H\<^sub>R
where "H\<^sub>R f \<equiv> \<lambda>\<mu>. \<mu> \<star> f"
(* TODO: Why do the following fail when I use @{thm ...} *)
text \<open>
Note that \<open>match_3\<close> and \<open>match_4\<close> are required for the next results.
\<close>
lemma endofunctor_H\<^sub>L:
assumes "ide g" and "g \<star> g \<noteq> null"
shows "endofunctor (Left g) (H\<^sub>L g)"
proof -
interpret L: subcategory V \<open>left g\<close> using assms left_hom_is_subcategory by simp
have *: "\<And>\<mu>. L.arr \<mu> \<Longrightarrow> H\<^sub>L g \<mu> = g \<star> \<mu>"
using assms H\<^sub>L_def by simp
have preserves_arr: "\<And>\<mu>. L.arr \<mu> \<Longrightarrow> L.arr (H\<^sub>L g \<mu>)"
- using assms * L.arr_char left_def match_4 by force
+ using assms * L.arr_char\<^sub>S\<^sub>b\<^sub>C 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
+ using H\<^sub>L_def L.arr_char\<^sub>S\<^sub>b\<^sub>C 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)
+ by (metis L.comp_def L.comp_simp L.seq_char\<^sub>S\<^sub>b\<^sub>C hcomp_simps\<^sub>W\<^sub>C(1) whisker_left preserves_arr)
qed
lemma endofunctor_H\<^sub>R:
assumes "ide f" and "f \<star> f \<noteq> null"
shows "endofunctor (Right f) (H\<^sub>R f)"
proof -
interpret R: subcategory V \<open>right f\<close> using assms right_hom_is_subcategory by simp
have *: "\<And>\<mu>. R.arr \<mu> \<Longrightarrow> H\<^sub>R f \<mu> = \<mu> \<star> f"
using assms H\<^sub>R_def by simp
have preserves_arr: "\<And>\<mu>. R.arr \<mu> \<Longrightarrow> R.arr (H\<^sub>R f \<mu>)"
- using assms * R.arr_char right_def match_3 by force
+ using assms * R.arr_char\<^sub>S\<^sub>b\<^sub>C 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
+ using H\<^sub>R_def R.arr_char\<^sub>S\<^sub>b\<^sub>C 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)
+ by (metis R.comp_def R.comp_simp R.seq_char\<^sub>S\<^sub>b\<^sub>C hcomp_simps\<^sub>W\<^sub>C(1) whisker_right preserves_arr)
qed
end
locale left_hom =
weak_composition V H +
S: subcategory V \<open>left \<omega>\<close>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a comp" (infixr "\<star>" 53)
and \<omega> :: 'a +
assumes arr_\<omega>: "arr \<omega>"
begin
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation in_hom ("\<guillemotleft>_ : _ \<Rightarrow> _\<guillemotright>")
notation S.comp (infixr "\<cdot>\<^sub>S" 55)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
lemma right_hcomp_closed:
assumes "\<guillemotleft>\<mu> : x \<Rightarrow>\<^sub>S y\<guillemotright>" and "\<guillemotleft>\<nu> : c \<Rightarrow> d\<guillemotright>" and "\<mu> \<star> \<nu> \<noteq> null"
shows "\<guillemotleft>\<mu> \<star> \<nu> : x \<star> c \<Rightarrow>\<^sub>S y \<star> d\<guillemotright>"
- using assms arr_\<omega> S.arr_char S.dom_simp S.cod_simp left_def match_4
+ using assms arr_\<omega> S.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<nu> \<mu>" and "S.seq \<tau> \<sigma>" and "\<mu> \<star> \<sigma> \<noteq> null"
shows "(\<nu> \<cdot>\<^sub>S \<mu>) \<star> (\<tau> \<cdot>\<^sub>S \<sigma>) = (\<nu> \<star> \<tau>) \<cdot>\<^sub>S (\<mu> \<star> \<sigma>)"
proof -
have 1: "\<nu> \<star> \<tau> \<noteq> null"
using assms hom_connected(1) [of \<nu> \<sigma>] hom_connected(2) [of \<nu> \<tau>] hom_connected(3-4)
S.dom_simp S.cod_simp
by force
have "(\<nu> \<cdot>\<^sub>S \<mu>) \<star> (\<tau> \<cdot>\<^sub>S \<sigma>) = (\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>)"
- using assms S.comp_char S.seq_char by metis
+ using assms S.comp_char S.seq_char\<^sub>S\<^sub>b\<^sub>C by metis
also have "... = (\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)"
- using assms interchange S.seq_char S.arr_char by simp
+ using assms interchange S.seq_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = (\<nu> \<star> \<tau>) \<cdot>\<^sub>S (\<mu> \<star> \<sigma>)"
using assms 1
- by (metis S.arr_char S.comp_char S.seq_char ext match_4 left_def)
+ by (metis S.arr_char\<^sub>S\<^sub>b\<^sub>C S.comp_char S.seq_char\<^sub>S\<^sub>b\<^sub>C ext match_4 left_def)
finally show ?thesis by blast
qed
lemma inv_char:
assumes "S.arr \<phi>" and "iso \<phi>"
shows "S.inverse_arrows \<phi> (inv \<phi>)"
and "S.inv \<phi> = inv \<phi>"
proof -
have 1: "S.arr (inv \<phi>)"
- using assms S.arr_char left_iff_left_inv
- by (intro S.arrI) meson
+ using assms S.arr_char\<^sub>S\<^sub>b\<^sub>C left_iff_left_inv
+ by (intro S.arrI\<^sub>S\<^sub>b\<^sub>C) meson
show "S.inv \<phi> = inv \<phi>"
- using assms 1 S.inv_char S.iso_char by blast
+ using assms 1 S.inv_char\<^sub>S\<^sub>b\<^sub>C S.iso_char\<^sub>S\<^sub>b\<^sub>C by blast
thus "S.inverse_arrows \<phi> (inv \<phi>)"
- using assms 1 S.iso_char S.inv_is_inverse by metis
+ using assms 1 S.iso_char\<^sub>S\<^sub>b\<^sub>C S.inv_is_inverse by metis
qed
lemma iso_char:
assumes "S.arr \<phi>"
shows "S.iso \<phi> \<longleftrightarrow> iso \<phi>"
- using assms S.iso_char inv_char by auto
+ using assms S.iso_char\<^sub>S\<^sub>b\<^sub>C inv_char by auto
end
locale right_hom =
weak_composition V H +
S: subcategory V \<open>right \<omega>\<close>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a comp" (infixr "\<star>" 53)
and \<omega> :: 'a +
assumes arr_\<omega>: "arr \<omega>"
begin
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation in_hom ("\<guillemotleft>_ : _ \<Rightarrow> _\<guillemotright>")
notation S.comp (infixr "\<cdot>\<^sub>S" 55)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
lemma left_hcomp_closed:
assumes "\<guillemotleft>\<mu> : x \<Rightarrow>\<^sub>S y\<guillemotright>" and "\<guillemotleft>\<nu> : c \<Rightarrow> d\<guillemotright>" and "\<nu> \<star> \<mu> \<noteq> null"
shows "\<guillemotleft>\<nu> \<star> \<mu> : c \<star> x \<Rightarrow>\<^sub>S d \<star> y\<guillemotright>"
- using assms arr_\<omega> S.arr_char S.dom_simp S.cod_simp right_def match_3
+ using assms arr_\<omega> S.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<nu> \<mu>" and "S.seq \<tau> \<sigma>" and "\<mu> \<star> \<sigma> \<noteq> null"
shows "(\<nu> \<cdot>\<^sub>S \<mu>) \<star> (\<tau> \<cdot>\<^sub>S \<sigma>) = (\<nu> \<star> \<tau>) \<cdot>\<^sub>S (\<mu> \<star> \<sigma>)"
proof -
have 1: "\<nu> \<star> \<tau> \<noteq> null"
using assms hom_connected(1) [of \<nu> \<sigma>] hom_connected(2) [of \<nu> \<tau>] hom_connected(3-4)
S.dom_simp S.cod_simp
by fastforce
have "(\<nu> \<cdot>\<^sub>S \<mu>) \<star> (\<tau> \<cdot>\<^sub>S \<sigma>) = (\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>)"
- using assms S.comp_char S.seq_char by metis
+ using assms S.comp_char S.seq_char\<^sub>S\<^sub>b\<^sub>C by metis
also have "... = (\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)"
- using assms interchange S.seq_char S.arr_char by simp
+ using assms interchange S.seq_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = (\<nu> \<star> \<tau>) \<cdot>\<^sub>S (\<mu> \<star> \<sigma>)"
using assms 1
- by (metis S.arr_char S.comp_char S.seq_char ext match_3 right_def)
+ by (metis S.arr_char\<^sub>S\<^sub>b\<^sub>C S.comp_char S.seq_char\<^sub>S\<^sub>b\<^sub>C ext match_3 right_def)
finally show ?thesis by blast
qed
lemma inv_char:
assumes "S.arr \<phi>" and "iso \<phi>"
shows "S.inverse_arrows \<phi> (inv \<phi>)"
and "S.inv \<phi> = inv \<phi>"
proof -
have 1: "S.arr (inv \<phi>)"
- using assms S.arr_char right_iff_right_inv
- by (intro S.arrI) meson
+ using assms S.arr_char\<^sub>S\<^sub>b\<^sub>C right_iff_right_inv
+ by (intro S.arrI\<^sub>S\<^sub>b\<^sub>C) meson
show "S.inv \<phi> = inv \<phi>"
- using assms 1 S.inv_char S.iso_char by blast
+ using assms 1 S.inv_char\<^sub>S\<^sub>b\<^sub>C S.iso_char\<^sub>S\<^sub>b\<^sub>C by blast
thus "S.inverse_arrows \<phi> (inv \<phi>)"
- using assms 1 S.iso_char S.inv_is_inverse by metis
+ using assms 1 S.iso_char\<^sub>S\<^sub>b\<^sub>C S.inv_is_inverse by metis
qed
lemma iso_char:
assumes "S.arr \<phi>"
shows "S.iso \<phi> \<longleftrightarrow> iso \<phi>"
- using assms S.iso_char inv_char by auto
+ using assms S.iso_char\<^sub>S\<^sub>b\<^sub>C inv_char by auto
end
subsection "Weak Units"
text \<open>
We now define a \emph{weak unit} to be an arrow \<open>a\<close> such that:
\begin{enumerate}
\item \<open>a \<star> a\<close> is isomorphic to \<open>a\<close>
(and hence \<open>a\<close> is a horizontally self-composable 1-cell).
\item Horizontal composition on the left with \<open>a\<close> is a fully faithful endofunctor of the
subcategory of arrows that are composable on the left with \<open>a\<close>.
\item Horizontal composition on the right with \<open>a\<close> is fully faithful endofunctor of the
subcategory of arrows that are composable on the right with \<open>a\<close>.
\end{enumerate}
\<close>
context weak_composition
begin
definition weak_unit :: "'a \<Rightarrow> bool"
where "weak_unit a \<equiv> a \<star> a \<cong> a \<and>
fully_faithful_functor (Left a) (Left a) (H\<^sub>L a) \<and>
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 \<star> a)" and "a \<star> a \<noteq> null"
proof -
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : a \<star> a \<Rightarrow> a\<guillemotright> \<and> iso \<phi>"
using assms weak_unit_def isomorphic_def by blast
have 1: "arr \<phi>" using \<phi> by blast
show "ide a" using \<phi> ide_cod by blast
thus "ide (a \<star> a)" using \<phi> ide_dom by force
thus "a \<star> a \<noteq> 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 "\<guillemotleft>a : a \<Rightarrow> a\<guillemotright>"
using assms weak_unit_self_composable left_def by auto
text \<open>
If \<open>a\<close> is a weak unit, then there exists a ``unit isomorphism'' \<open>\<guillemotleft>\<iota> : a \<star> a \<Rightarrow> a\<guillemotright>\<close>.
It need not be unique, but we may choose one arbitrarily.
\<close>
definition some_unit
where "some_unit a \<equiv> SOME \<iota>. iso \<iota> \<and> \<guillemotleft>\<iota> : a \<star> a \<Rightarrow> a\<guillemotright>"
lemma iso_some_unit:
assumes "weak_unit a"
shows "iso (some_unit a)"
and "\<guillemotleft>some_unit a : a \<star> a \<Rightarrow> a\<guillemotright>"
proof -
let ?P = "\<lambda>\<iota>. iso \<iota> \<and> \<guillemotleft>\<iota> : a \<star> a \<Rightarrow> a\<guillemotright>"
have "\<exists>\<iota>. ?P \<iota>"
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 "\<guillemotleft>some_unit a : a \<star> a \<Rightarrow> a\<guillemotright>" using 1 by blast
qed
text \<open>
The \emph{sources} of an arbitrary arrow \<open>\<mu>\<close> are the weak units that are composable with \<open>\<mu>\<close>
on the right. Similarly, the \emph{targets} of \<open>\<mu>\<close> are the weak units that are composable
with \<open>\<mu>\<close> on the left.
\<close>
definition sources
where "sources \<mu> \<equiv> {a. weak_unit a \<and> \<mu> \<star> a \<noteq> null}"
lemma sourcesI [intro]:
assumes "weak_unit a" and "\<mu> \<star> a \<noteq> null"
shows "a \<in> sources \<mu>"
using assms sources_def by blast
lemma sourcesD [dest]:
assumes "a \<in> sources \<mu>"
shows "ide a" and "weak_unit a" and "\<mu> \<star> a \<noteq> null"
using assms sources_def weak_unit_self_composable by auto
definition targets
where "targets \<mu> \<equiv> {b. weak_unit b \<and> b \<star> \<mu> \<noteq> null}"
lemma targetsI [intro]:
assumes "weak_unit b" and "b \<star> \<mu> \<noteq> null"
shows "b \<in> targets \<mu>"
using assms targets_def by blast
lemma targetsD [dest]:
assumes "b \<in> targets \<mu>"
shows "ide b" and "weak_unit b" and "b \<star> \<mu> \<noteq> null"
using assms targets_def weak_unit_self_composable by auto
lemma sources_dom [simp]:
assumes "arr \<mu>"
shows "sources (dom \<mu>) = sources \<mu>"
using assms hom_connected(1) by blast
lemma sources_cod [simp]:
assumes "arr \<mu>"
shows "sources (cod \<mu>) = sources \<mu>"
using assms hom_connected(3) by blast
lemma targets_dom [simp]:
assumes "arr \<mu>"
shows "targets (dom \<mu>) = targets \<mu>"
using assms hom_connected(2) by blast
lemma targets_cod [simp]:
assumes "arr \<mu>"
shows "targets (cod \<mu>) = targets \<mu>"
using assms hom_connected(4) by blast
lemma weak_unit_iff_self_source:
shows "weak_unit a \<longleftrightarrow> a \<in> sources a"
using weak_unit_self_composable by auto
lemma weak_unit_iff_self_target:
shows "weak_unit b \<longleftrightarrow> b \<in> targets b"
using weak_unit_self_composable by auto
abbreviation (input) in_hhom\<^sub>W\<^sub>C ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>W\<^sub>C _\<guillemotright>")
where "in_hhom\<^sub>W\<^sub>C \<mu> f g \<equiv> arr \<mu> \<and> f \<in> sources \<mu> \<and> g \<in> targets \<mu>"
lemma sources_hcomp:
assumes "\<nu> \<star> \<mu> \<noteq> null"
shows "sources (\<nu> \<star> \<mu>) = sources \<mu>"
using assms match_1 match_3 null_agreement by blast
lemma targets_hcomp:
assumes "\<nu> \<star> \<mu> \<noteq> null"
shows "targets (\<nu> \<star> \<mu>) = targets \<nu>"
using assms match_2 match_4 null_agreement by blast
lemma H\<^sub>R_preserved_along_iso:
assumes "weak_unit a" and "a \<cong> a'"
shows "endofunctor (Right a) (H\<^sub>R a')"
proof -
have a: "ide a \<and> 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 \<open>right a\<close> using a right_hom_is_subcategory by simp
have *: "\<And>\<mu>. R.arr \<mu> \<Longrightarrow> H\<^sub>R a' \<mu> = \<mu> \<star> a'"
using assms H\<^sub>R_def by simp
have preserves_arr: "\<And>\<mu>. R.arr \<mu> \<Longrightarrow> R.arr (H\<^sub>R a' \<mu>)"
- using assms a' * R.arr_char right_def weak_unit_def weak_unit_self_composable
+ using assms a' * R.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "\<And>\<mu>. \<not> R.arr \<mu> \<Longrightarrow> H\<^sub>R a' \<mu> = R.null"
- using assms R.arr_char R.null_char right_def H\<^sub>R_def null_agreement
+ using assms R.arr_char\<^sub>S\<^sub>b\<^sub>C R.null_char right_def H\<^sub>R_def null_agreement
right_respects_isomorphic
by metis
fix \<mu>
assume "R.arr \<mu>"
hence \<mu>: "R.arr \<mu> \<and> arr \<mu> \<and> right a \<mu> \<and> right a' \<mu> \<and> \<mu> \<star> a \<noteq> null \<and> \<mu> \<star> a' \<noteq> null"
- using assms R.arr_char right_respects_isomorphic composable_implies_arr null_agreement
+ using assms R.arr_char\<^sub>S\<^sub>b\<^sub>C right_respects_isomorphic composable_implies_arr null_agreement
right_def
by metis
show "R.arr (H\<^sub>R a' \<mu>)" using \<mu> preserves_arr by blast
show "R.dom (H\<^sub>R a' \<mu>) = H\<^sub>R a' (R.dom \<mu>)"
- using a' \<mu> * R.arr_char R.dom_char preserves_arr hom_connected(1) right_def
+ using a' \<mu> * R.arr_char\<^sub>S\<^sub>b\<^sub>C R.dom_char\<^sub>S\<^sub>b\<^sub>C preserves_arr hom_connected(1) right_def
by simp
show "R.cod (H\<^sub>R a' \<mu>) = H\<^sub>R a' (R.cod \<mu>)"
- using a' \<mu> * R.arr_char R.cod_char preserves_arr hom_connected(3) right_def
+ using a' \<mu> * R.arr_char\<^sub>S\<^sub>b\<^sub>C R.cod_char\<^sub>S\<^sub>b\<^sub>C preserves_arr hom_connected(3) right_def
by simp
next
fix \<mu> \<nu>
assume \<mu>\<nu>: "R.seq \<nu> \<mu>"
have \<mu>: "R.arr \<mu> \<and> arr \<mu> \<and> right a \<mu> \<and> right a' \<mu> \<and> \<mu> \<star> a \<noteq> null \<and> \<mu> \<star> a' \<noteq> null"
- using assms \<mu>\<nu> R.arr_char right_respects_isomorphic composable_implies_arr
+ using assms \<mu>\<nu> R.arr_char\<^sub>S\<^sub>b\<^sub>C right_respects_isomorphic composable_implies_arr
null_agreement right_def
by (elim R.seqE) metis
have \<nu>: "\<guillemotleft>\<nu> : R.cod \<mu> \<rightarrow> R.cod \<nu>\<guillemotright> \<and> arr \<nu> \<and>
right a \<nu> \<and> H \<nu> a \<noteq> null \<and> right a' \<nu> \<and> H \<nu> a' \<noteq> null"
by (metis "*" R.cod_simp R.comp_def R.inclusion R.not_arr_null R.null_char R.seqE
\<mu>\<nu> in_homI preserves_arr right_def)
show "H\<^sub>R a' (R.comp \<nu> \<mu>) = R.comp (H\<^sub>R a' \<nu>) (H\<^sub>R a' \<mu>)"
proof -
have 1: "R.arr (H\<^sub>R a' \<nu>)"
using \<nu> preserves_arr by blast
have 2: "seq (\<nu> \<star> a') (\<mu> \<star> a')"
- using a' \<mu> \<nu> R.arr_char R.inclusion R.dom_char R.cod_char
+ using a' \<mu> \<nu> R.arr_char\<^sub>S\<^sub>b\<^sub>C R.inclusion R.dom_char\<^sub>S\<^sub>b\<^sub>C R.cod_char\<^sub>S\<^sub>b\<^sub>C
isomorphic_implies_equicomposable
by auto
show ?thesis
using a' \<mu> \<nu> \<mu>\<nu> 1 2 preserves_arr H\<^sub>R_def R.dom_simp R.cod_simp R.comp_char
- R.seq_char R.inclusion whisker_right
+ R.seq_char\<^sub>S\<^sub>b\<^sub>C R.inclusion whisker_right
by metis
qed
qed
qed
lemma H\<^sub>L_preserved_along_iso:
assumes "weak_unit a" and "a \<cong> a'"
shows "endofunctor (Left a) (H\<^sub>L a')"
proof -
have a: "ide a \<and> 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 \<open>left a\<close> using a left_hom_is_subcategory by simp
have *: "\<And>\<mu>. L.arr \<mu> \<Longrightarrow> H\<^sub>L a' \<mu> = a' \<star> \<mu>"
using assms H\<^sub>L_def by simp
have preserves_arr: "\<And>\<mu>. L.arr \<mu> \<Longrightarrow> L.arr (H\<^sub>L a' \<mu>)"
- using assms a' * L.arr_char left_def weak_unit_def weak_unit_self_composable
+ using assms a' * L.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "\<And>\<mu>. \<not> L.arr \<mu> \<Longrightarrow> H\<^sub>L a' \<mu> = L.null"
- using assms L.arr_char L.null_char left_def H\<^sub>L_def null_agreement
+ using assms L.arr_char\<^sub>S\<^sub>b\<^sub>C L.null_char left_def H\<^sub>L_def null_agreement
left_respects_isomorphic
by metis
fix \<mu>
assume "L.arr \<mu>"
hence \<mu>: "L.arr \<mu> \<and> arr \<mu> \<and> left a \<mu> \<and> left a' \<mu> \<and> a \<star> \<mu> \<noteq> null \<and> a' \<star> \<mu> \<noteq> null"
- using assms L.arr_char left_respects_isomorphic composable_implies_arr null_agreement
+ using assms L.arr_char\<^sub>S\<^sub>b\<^sub>C left_respects_isomorphic composable_implies_arr null_agreement
left_def
by metis
show "L.arr (H\<^sub>L a' \<mu>)" using \<mu> preserves_arr by blast
show "L.dom (H\<^sub>L a' \<mu>) = H\<^sub>L a' (L.dom \<mu>)"
- using a' \<mu> * L.arr_char L.dom_char preserves_arr hom_connected(2) left_def
+ using a' \<mu> * L.arr_char\<^sub>S\<^sub>b\<^sub>C L.dom_char\<^sub>S\<^sub>b\<^sub>C preserves_arr hom_connected(2) left_def
by simp
show "L.cod (H\<^sub>L a' \<mu>) = H\<^sub>L a' (L.cod \<mu>)"
- using a' \<mu> * L.arr_char L.cod_char preserves_arr hom_connected(4) left_def
+ using a' \<mu> * L.arr_char\<^sub>S\<^sub>b\<^sub>C L.cod_char\<^sub>S\<^sub>b\<^sub>C preserves_arr hom_connected(4) left_def
by simp
next
fix \<mu> \<nu>
assume \<mu>\<nu>: "L.seq \<nu> \<mu>"
have "L.arr \<mu>"
using \<mu>\<nu> by (elim L.seqE, auto)
hence \<mu>: "L.arr \<mu> \<and> arr \<mu> \<and> left a \<mu> \<and> left a' \<mu> \<and> a \<star> \<mu> \<noteq> null \<and> a' \<star> \<mu> \<noteq> null"
- using assms L.arr_char left_respects_isomorphic composable_implies_arr null_agreement
+ using assms L.arr_char\<^sub>S\<^sub>b\<^sub>C left_respects_isomorphic composable_implies_arr null_agreement
left_def
by metis
have \<nu>: "\<guillemotleft>\<nu> : L.cod \<mu> \<Rightarrow> L.cod \<nu>\<guillemotright> \<and> arr \<nu> \<and>
left a \<nu> \<and> a \<star> \<nu> \<noteq> null \<and> left a' \<nu> \<and> a' \<star> \<nu> \<noteq> null"
- by (metis (mono_tags, opaque_lifting) L.arrE L.cod_simp L.seq_char \<mu>\<nu> assms(2)
+ by (metis (mono_tags, opaque_lifting) L.arrE L.cod_simp L.seq_char\<^sub>S\<^sub>b\<^sub>C \<mu>\<nu> assms(2)
in_homI seqE left_def left_respects_isomorphic)
show "H\<^sub>L a' (L.comp \<nu> \<mu>) = L.comp (H\<^sub>L a' \<nu>) (H\<^sub>L a' \<mu>)"
proof -
have 1: "L.arr (H\<^sub>L a' \<nu>)"
using \<nu> preserves_arr by blast
have 2: "seq (a' \<star> \<nu>) (a' \<star> \<mu>)"
- using a' \<mu> \<nu> L.arr_char L.inclusion L.dom_char L.cod_char
+ using a' \<mu> \<nu> L.arr_char\<^sub>S\<^sub>b\<^sub>C L.inclusion L.dom_char\<^sub>S\<^sub>b\<^sub>C L.cod_char\<^sub>S\<^sub>b\<^sub>C
isomorphic_implies_equicomposable
by auto
show ?thesis
using a' \<mu> \<nu> \<mu>\<nu> 1 2 preserves_arr H\<^sub>L_def L.dom_simp L.cod_simp L.comp_char
- L.seq_char L.inclusion whisker_left
+ L.seq_char\<^sub>S\<^sub>b\<^sub>C L.inclusion whisker_left
by metis
qed
qed
qed
end
subsection "Regularity"
text \<open>
We call a weak composition \emph{regular} if \<open>f \<star> a \<cong> f\<close> whenever \<open>a\<close> is a source of
1-cell \<open>f\<close>, and \<open>b \<star> f \<cong> f\<close> whenever \<open>b\<close> is a target of \<open>f\<close>. A consequence of regularity
is that horizontal composability of 2-cells is fully determined by their sets of
sources and targets.
\<close>
locale regular_weak_composition =
weak_composition +
assumes comp_ide_source: "\<lbrakk> a \<in> sources f; ide f \<rbrakk> \<Longrightarrow> f \<star> a \<cong> f"
and comp_target_ide: "\<lbrakk> b \<in> targets f; ide f \<rbrakk> \<Longrightarrow> b \<star> f \<cong> f"
begin
lemma sources_determine_composability:
assumes "a \<in> sources \<tau>"
shows "\<tau> \<star> \<mu> \<noteq> null \<longleftrightarrow> a \<star> \<mu> \<noteq> null"
proof -
have *: "\<And>\<tau>. ide \<tau> \<and> a \<in> sources \<tau> \<Longrightarrow> \<tau> \<star> \<mu> \<noteq> null \<longleftrightarrow> a \<star> \<mu> \<noteq> null"
using assms comp_ide_source isomorphic_implies_equicomposable match_1 match_3
by (meson sourcesD(3))
show ?thesis
proof -
have "arr \<tau>" using assms composable_implies_arr by auto
thus ?thesis
using assms * [of "dom \<tau>"] hom_connected(1) by auto
qed
qed
lemma targets_determine_composability:
assumes "b \<in> targets \<mu>"
shows "\<tau> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<tau> \<star> b \<noteq> null"
proof -
have *: "\<And>\<mu>. ide \<mu> \<and> b \<in> targets \<mu> \<Longrightarrow> \<tau> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<tau> \<star> b \<noteq> null"
using assms comp_target_ide isomorphic_implies_equicomposable match_2 match_4
by (meson targetsD(3))
show ?thesis
proof -
have "arr \<mu>" using assms composable_implies_arr by auto
thus ?thesis
using assms * [of "dom \<mu>"] hom_connected(2) by auto
qed
qed
lemma composable_if_connected:
assumes "sources \<nu> \<inter> targets \<mu> \<noteq> {}"
shows "\<nu> \<star> \<mu> \<noteq> null"
using assms targets_determine_composability by blast
lemma connected_if_composable:
assumes "\<nu> \<star> \<mu> \<noteq> null"
shows "sources \<nu> = targets \<mu>"
using assms sources_determine_composability targets_determine_composability by blast
lemma iso_hcomp\<^sub>R\<^sub>W\<^sub>C:
assumes "iso \<mu>" and "iso \<nu>" and "sources \<nu> \<inter> targets \<mu> \<noteq> {}"
shows "iso (\<nu> \<star> \<mu>)"
and "inverse_arrows (\<nu> \<star> \<mu>) (inv \<nu> \<star> inv \<mu>)"
proof -
have \<mu>: "arr \<mu> \<and> \<guillemotleft>\<mu> : dom \<mu> \<Rightarrow> cod \<mu>\<guillemotright> \<and>
iso \<mu> \<and> \<guillemotleft>inv \<mu> : cod \<mu> \<Rightarrow> dom \<mu>\<guillemotright>"
using assms inv_in_hom arr_iff_in_hom iso_is_arr by auto
have \<nu>: "arr \<nu> \<and> \<guillemotleft>\<nu> : dom \<nu> \<Rightarrow> cod \<nu>\<guillemotright> \<and>
iso \<nu> \<and> \<guillemotleft>inv \<nu> : cod \<nu> \<Rightarrow> dom \<nu>\<guillemotright>"
using assms inv_in_hom by blast
have 1: "sources (inv \<nu>) \<inter> targets (inv \<mu>) \<noteq> {}"
using assms \<mu> \<nu> sources_dom sources_cod targets_dom targets_cod arr_inv cod_inv
by metis
show "inverse_arrows (\<nu> \<star> \<mu>) (inv \<nu> \<star> inv \<mu>)"
using assms 1 \<mu> \<nu> inv_in_hom inv_is_inverse comp_inv_arr
interchange [of "inv \<nu>" \<nu> "inv \<mu>" \<mu>] composable_if_connected
ide_hcomp\<^sub>W\<^sub>C sources_dom targets_dom interchange [of \<nu> "inv \<nu>" \<mu> "inv \<mu>"]
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 (\<nu> \<star> \<mu>)" by auto
qed
lemma inv_hcomp\<^sub>R\<^sub>W\<^sub>C:
assumes "iso \<mu>" and "iso \<nu>" and "sources \<nu> \<inter> targets \<mu> \<noteq> {}"
shows "inv (\<nu> \<star> \<mu>) = inv \<nu> \<star> inv \<mu>"
using assms iso_hcomp\<^sub>R\<^sub>W\<^sub>C(2) [of \<mu> \<nu>] inverse_arrow_unique [of "H \<nu> \<mu>"] inv_is_inverse
by auto
end
subsection "Associativity"
text \<open>
An \emph{associative weak composition} consists of a weak composition that has been
equipped with an \emph{associator} isomorphism: \<open>\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>\<close>
for each composable triple \<open>(f, g, h)\<close> of 1-cells, subject to naturality and
coherence conditions.
\<close>
locale associative_weak_composition =
weak_composition +
fixes \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
assumes assoc_in_vhom\<^sub>A\<^sub>W\<^sub>C:
"\<lbrakk> ide f; ide g; ide h; f \<star> g \<noteq> null; g \<star> h \<noteq> null \<rbrakk> \<Longrightarrow>
\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
and assoc_naturality\<^sub>A\<^sub>W\<^sub>C:
"\<lbrakk> \<tau> \<star> \<mu> \<noteq> null; \<mu> \<star> \<nu> \<noteq> null \<rbrakk> \<Longrightarrow>
\<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>) = (\<tau> \<star> \<mu> \<star> \<nu>) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>]"
and iso_assoc\<^sub>A\<^sub>W\<^sub>C: "\<lbrakk> ide f; ide g; ide h; f \<star> g \<noteq> null; g \<star> h \<noteq> null \<rbrakk> \<Longrightarrow> iso \<a>[f, g, h]"
and pentagon\<^sub>A\<^sub>W\<^sub>C:
"\<lbrakk> ide f; ide g; ide h; ide k; sources f \<inter> targets g \<noteq> {};
sources g \<inter> targets h \<noteq> {}; sources h \<inter> targets k \<noteq> {} \<rbrakk> \<Longrightarrow>
(f \<star> \<a>[g, h, k]) \<cdot> \<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k) = \<a>[f, g, h \<star> k] \<cdot> \<a>[f \<star> 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 \<star> g \<noteq> null" and "g \<star> h \<noteq> null"
shows "sources \<a>[f, g, h] = sources h" and "targets \<a>[f, g, h] = targets f"
and "\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
using assms assoc_in_vhom\<^sub>A\<^sub>W\<^sub>C by simp
show "sources \<a>[f, g, h] = sources h"
using assms 1 sources_dom [of "\<a>[f, g, h]"] sources_hcomp match_3
by (elim in_homE, auto)
show "targets \<a>[f, g, h] = targets f"
using assms 1 targets_cod [of "\<a>[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 \<star> g \<noteq> null" and "g \<star> h \<noteq> null"
shows "arr \<a>[f, g, h]"
and "dom \<a>[f, g, h] = (f \<star> g) \<star> h"
and "cod \<a>[f, g, h] = f \<star> g \<star> h"
proof -
have 1: "\<guillemotleft>\<a>[f, g, h] : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
using assms assoc_in_hom\<^sub>A\<^sub>W\<^sub>C by auto
show "arr \<a>[f, g, h]" using 1 by auto
show "dom \<a>[f, g, h] = (f \<star> g) \<star> h" using 1 by auto
show "cod \<a>[f, g, h] = f \<star> g \<star> 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 \<star> g \<noteq> null" and "g \<star> h \<noteq> null"
shows "sources (inv \<a>[f, g, h]) = sources h" and "targets (inv \<a>[f, g, h]) = targets f"
and "\<guillemotleft>inv \<a>[f, g, h] : f \<star> g \<star> h \<Rightarrow> (f \<star> g) \<star> h\<guillemotright>"
proof -
show 1: "\<guillemotleft>inv \<a>[f, g, h] : f \<star> g \<star> h \<Rightarrow> (f \<star> g) \<star> h\<guillemotright>"
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 \<a>[f, g, h]) = sources h"
using assms 1 sources_hcomp [of "f \<star> g" h] sources_cod match_3 null_agreement
by (elim in_homE, metis)
show "targets (inv \<a>[f, g, h]) = targets f"
using assms 1 targets_hcomp [of f "g \<star> 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 \<star> g \<noteq> null" and "g \<star> h \<noteq> null"
shows "arr (inv \<a>[f, g, h])"
and "dom (inv \<a>[f, g, h]) = f \<star> g \<star> h"
and "cod (inv \<a>[f, g, h]) = (f \<star> g) \<star> h"
proof -
have 1: "\<guillemotleft>inv \<a>[f, g, h] : f \<star> g \<star> h \<Rightarrow> (f \<star> g) \<star> h\<guillemotright>"
using assms assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C by auto
show "arr (inv \<a>[f, g, h])" using 1 by auto
show "dom (inv \<a>[f, g, h]) = f \<star> g \<star> h" using 1 by auto
show "cod (inv \<a>[f, g, h]) = (f \<star> g) \<star> h" using 1 by auto
qed
lemma assoc'_naturality\<^sub>A\<^sub>W\<^sub>C:
assumes "\<tau> \<star> \<mu> \<noteq> null" and "\<mu> \<star> \<nu> \<noteq> null"
shows "inv \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> (\<tau> \<star> \<mu> \<star> \<nu>) = ((\<tau> \<star> \<mu>) \<star> \<nu>) \<cdot> inv \<a>[dom \<tau>, dom \<mu>, dom \<nu>]"
proof -
have \<tau>\<mu>\<nu>: "arr \<tau> \<and> arr \<mu> \<and> arr \<nu>"
using assms composable_implies_arr by simp
have 0: "dom \<tau> \<star> dom \<mu> \<noteq> null \<and> dom \<mu> \<star> dom \<nu> \<noteq> null \<and>
cod \<tau> \<star> cod \<mu> \<noteq> null \<and> cod \<mu> \<star> cod \<nu> \<noteq> null"
using assms \<tau>\<mu>\<nu> hom_connected by simp
have 1: "\<guillemotleft>\<tau> \<star> \<mu> \<star> \<nu> : dom \<tau> \<star> dom \<mu> \<star> dom \<nu> \<Rightarrow> cod \<tau> \<star> cod \<mu> \<star> cod \<nu>\<guillemotright>"
using assms match_4 by auto
have 2: "\<guillemotleft>(\<tau> \<star> \<mu>) \<star> \<nu> : (dom \<tau> \<star> dom \<mu>) \<star> dom \<nu> \<Rightarrow> (cod \<tau> \<star> cod \<mu>) \<star> cod \<nu>\<guillemotright>"
using assms match_3 by auto
have "(inv \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> (\<tau> \<star> \<mu> \<star> \<nu>)) \<cdot> \<a>[dom \<tau>, dom \<mu>, dom \<nu>] = (\<tau> \<star> \<mu>) \<star> \<nu>"
proof -
have "(\<tau> \<star> \<mu>) \<star> \<nu> = (inv \<a>[cod \<tau>, cod \<mu>, cod \<nu>] \<cdot> \<a>[cod \<tau>, cod \<mu>, cod \<nu>]) \<cdot> ((\<tau> \<star> \<mu>) \<star> \<nu>)"
using 0 2 \<tau>\<mu>\<nu> 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 \<tau>\<mu>\<nu> 0 2 assoc_naturality\<^sub>A\<^sub>W\<^sub>C comp_assoc by metis
qed
thus ?thesis
using 0 1 2 \<tau>\<mu>\<nu> 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 \<open>
For an associative weak composition with a chosen unit isomorphism \<open>\<iota> : a \<star> a \<Rightarrow> a\<close>,
where \<open>a\<close> is a weak unit, horizontal composition on the right by \<open>a\<close> is a fully faithful
endofunctor \<open>R\<close> of the subcategory of arrows composable on the right with \<open>a\<close>, and is
consequently an endo-equivalence of that subcategory. This equivalence, together with the
associator isomorphisms and unit isomorphism \<open>\<iota>\<close>, canonically associate, with each
identity arrow \<open>f\<close> composable on the right with \<open>a\<close>, a \emph{right unit} isomorphism
\<open>\<guillemotleft>\<r>[f] : f \<star> a \<Rightarrow> f\<guillemotright>\<close>. These isomorphisms are the components of a natural isomorphism
from \<open>R\<close> to the identity functor.
\<close>
locale right_hom_with_unit =
associative_weak_composition V H \<a> +
right_hom V H a
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a comp" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<iota> :: 'a
and a :: 'a +
assumes weak_unit_a: "weak_unit a"
and \<iota>_in_hom: "\<guillemotleft>\<iota> : a \<star> a \<Rightarrow> a\<guillemotright>"
and iso_\<iota>: "iso \<iota>"
begin
abbreviation R
where "R \<equiv> 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 ("\<r>[_]")
where "runit f \<equiv> THE \<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow>\<^sub>S f\<guillemotright> \<and> R \<mu> = (f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a]"
lemma iso_unit:
shows "S.iso \<iota>" and "\<guillemotleft>\<iota> : a \<star> a \<Rightarrow>\<^sub>S a\<guillemotright>"
proof -
show "\<guillemotleft>\<iota> : a \<star> a \<Rightarrow>\<^sub>S a\<guillemotright>"
- 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
+ using weak_unit_a S.ide_char S.arr_char\<^sub>S\<^sub>b\<^sub>C right_def weak_unit_self_composable
+ S.ideD(1) R.preserves_arr H\<^sub>R_def S.in_hom_char\<^sub>S\<^sub>b\<^sub>C right_def
\<iota>_in_hom S.ideD(1) hom_connected(3) in_homE
by metis
thus "S.iso \<iota>"
using iso_\<iota> iso_char by blast
qed
lemma characteristic_iso:
assumes "S.ide f"
shows "\<guillemotleft>\<a>[f, a, a] : (f \<star> a) \<star> a \<Rightarrow>\<^sub>S f \<star> a \<star> a\<guillemotright>"
and "\<guillemotleft>f \<star> \<iota> : f \<star> a \<star> a \<Rightarrow>\<^sub>S f \<star> a\<guillemotright>"
and "\<guillemotleft>(f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a] : R (R f) \<Rightarrow>\<^sub>S R f\<guillemotright>"
and "S.iso ((f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a])"
proof -
have f: "S.ide f \<and> ide f"
- using assms S.ide_char by simp
+ using assms S.ide_char\<^sub>S\<^sub>b\<^sub>C by simp
have a: "weak_unit a \<and> ide a \<and> S.ide a"
- using weak_unit_a S.ide_char weak_unit_def S.arr_char right_def
+ using weak_unit_a S.ide_char\<^sub>S\<^sub>b\<^sub>C weak_unit_def S.arr_char\<^sub>S\<^sub>b\<^sub>C right_def
weak_unit_self_composable
by metis
have fa: "f \<star> a \<noteq> null \<and> (f \<star> a) \<star> a \<noteq> null \<and> ((f \<star> a) \<star> a) \<star> a \<noteq> 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 \<star> a \<noteq> 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 \<star> \<iota> \<noteq> null"
- using assms S.ide_char right_def S.arr_char hom_connected(4) \<iota>_in_hom by auto
+ using assms S.ide_char right_def S.arr_char\<^sub>S\<^sub>b\<^sub>C hom_connected(4) \<iota>_in_hom by auto
show assoc_in_hom: "\<guillemotleft>\<a>[f, a, a] : (f \<star> a) \<star> a \<Rightarrow>\<^sub>S f \<star> a \<star> a\<guillemotright>"
- using a f fa hom_connected(1) [of "\<a>[f, a, a]" a] S.arr_char right_def
- match_3 match_4 S.in_hom_char weak_unit_self_composable
+ using a f fa hom_connected(1) [of "\<a>[f, a, a]" a] S.arr_char\<^sub>S\<^sub>b\<^sub>C right_def
+ match_3 match_4 S.in_hom_char\<^sub>S\<^sub>b\<^sub>C weak_unit_self_composable
by auto
show 1: "\<guillemotleft>f \<star> \<iota> : f \<star> a \<star> a \<Rightarrow>\<^sub>S f \<star> a\<guillemotright>"
using a f fa iso_unit left_hcomp_closed
by (simp add: f_ia ide_in_hom)
have unit_part: "\<guillemotleft>f \<star> \<iota> : f \<star> a \<star> a \<Rightarrow>\<^sub>S f \<star> a\<guillemotright> \<and> S.iso (f \<star> \<iota>)"
proof -
have "S.iso (f \<star> \<iota>)"
- using a f fa f_ia 1 VoV.arr_char VxV.inv_simp
- inv_in_hom hom_connected(2) [of f "inv \<iota>"] VoV.arr_char VoV.iso_char
+ using a f fa f_ia 1 VoV.arr_char\<^sub>S\<^sub>b\<^sub>C VxV.inv_simp
+ inv_in_hom hom_connected(2) [of f "inv \<iota>"] VoV.arr_char\<^sub>S\<^sub>b\<^sub>C VoV.iso_char\<^sub>S\<^sub>b\<^sub>C
preserves_iso iso_char iso_\<iota> S.dom_simp S.cod_simp
by auto
thus ?thesis using 1 by blast
qed
show "S.iso ((f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a])"
using assms a f fa aa hom_connected(1) [of "\<a>[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
+ iso_assoc\<^sub>A\<^sub>W\<^sub>C iso_char S.arr_char\<^sub>S\<^sub>b\<^sub>C unit_part assoc_in_hom isos_compose
S.isos_compose S.seqI' by auto
show "\<guillemotleft>(f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a] : R (R f) \<Rightarrow>\<^sub>S R f\<guillemotright>"
unfolding H\<^sub>R_def using unit_part assoc_in_hom by blast
qed
lemma runit_char:
assumes "S.ide f"
shows "\<guillemotleft>\<r>[f] : R f \<Rightarrow>\<^sub>S f\<guillemotright>" and "R \<r>[f] = (f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow>\<^sub>S f\<guillemotright> \<and> R \<mu> = (f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a]"
proof -
let ?P = "\<lambda>\<mu>. \<guillemotleft>\<mu> : R f \<Rightarrow>\<^sub>S f\<guillemotright> \<and> R \<mu> = (f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a]"
show "\<exists>!\<mu>. ?P \<mu>"
proof -
have "\<exists>\<mu>. ?P \<mu>"
- using assms S.ide_char S.arr_char R.preserves_ide characteristic_iso(3) R.is_full
+ using assms S.ide_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C R.preserves_ide characteristic_iso(3) R.is_full
by auto
moreover have "\<forall>\<mu> \<mu>'. ?P \<mu> \<and> ?P \<mu>' \<longrightarrow> \<mu> = \<mu>'"
using R.is_faithful S.in_homE by metis
ultimately show ?thesis by blast
qed
hence "?P (THE \<mu>. ?P \<mu>)"
using theI' [of ?P] by fastforce
hence 1: "?P \<r>[f]"
unfolding runit_def by blast
show "\<guillemotleft>\<r>[f] : R f \<Rightarrow>\<^sub>S f\<guillemotright>" using 1 by fast
show "R \<r>[f] = (f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a]" using 1 by fast
qed
lemma iso_runit:
assumes "S.ide f"
shows "S.iso \<r>[f]"
using assms characteristic_iso(4) runit_char R.reflects_iso by metis
lemma runit_eqI:
assumes "\<guillemotleft>f : a \<Rightarrow>\<^sub>S b\<guillemotright>" and "\<guillemotleft>\<mu> : R f \<Rightarrow>\<^sub>S f\<guillemotright>"
and "R \<mu> = ((f \<star> \<iota>) \<cdot>\<^sub>S \<a>[f, a, a])"
shows "\<mu> = \<r>[f]"
using assms S.ide_cod runit_char [of f] by blast
lemma runit_naturality:
assumes "S.arr \<mu>"
shows "\<r>[S.cod \<mu>] \<cdot>\<^sub>S R \<mu> = \<mu> \<cdot>\<^sub>S \<r>[S.dom \<mu>]"
proof -
have 1: "\<guillemotleft>\<r>[S.cod \<mu>] \<cdot>\<^sub>S R \<mu> : R (S.dom \<mu>) \<Rightarrow>\<^sub>S S.cod \<mu>\<guillemotright>"
using assms runit_char(1) S.ide_cod by blast
have 2: "S.par (\<r>[S.cod \<mu>] \<cdot>\<^sub>S R \<mu>) (\<mu> \<cdot>\<^sub>S \<r>[S.dom \<mu>])"
using assms 1 S.ide_dom runit_char(1)
by (metis S.comp_in_homI' S.in_homE)
moreover have "R (\<r>[S.cod \<mu>] \<cdot>\<^sub>S R \<mu>) = R (\<mu> \<cdot>\<^sub>S \<r>[S.dom \<mu>])"
proof -
have 3: "\<guillemotleft>\<mu> \<star> a \<star> a : S.dom \<mu> \<star> a \<star> a \<Rightarrow>\<^sub>S S.cod \<mu> \<star> a \<star> a\<guillemotright>"
- 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
+ using assms weak_unit_a R.preserves_hom H\<^sub>R_def S.arr_iff_in_hom S.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis match_4 weak_unit_in_vhom weak_unit_self_right S.in_hom_char\<^sub>S\<^sub>b\<^sub>C
left_hcomp_closed S.not_arr_null S.null_char)
have 4: "R (\<r>[S.cod \<mu>] \<cdot>\<^sub>S R \<mu>) = R \<r>[S.cod \<mu>] \<cdot>\<^sub>S R (R \<mu>)"
using assms 1 R.as_nat_trans.preserves_comp_2 by blast
also have 5: "... = ((S.cod \<mu> \<star> \<iota>) \<cdot>\<^sub>S \<a>[S.cod \<mu>, a, a]) \<cdot>\<^sub>S ((\<mu> \<star> a) \<star> a)"
using assms R.preserves_arr runit_char S.ide_cod H\<^sub>R_def by auto
also have 6: "... = (S.cod \<mu> \<star> \<iota>) \<cdot>\<^sub>S \<a>[S.cod \<mu>, a, a] \<cdot>\<^sub>S ((\<mu> \<star> a) \<star> a)"
using assms S.comp_assoc by simp
also have "... = (S.cod \<mu> \<star> \<iota>) \<cdot>\<^sub>S (\<mu> \<star> a \<star> a) \<cdot>\<^sub>S \<a>[S.dom \<mu>, a, a]"
proof -
have "(\<mu> \<star> a \<star> a) \<cdot>\<^sub>S \<a>[S.dom \<mu>, a, a] = \<a>[S.cod \<mu>, a, a] \<cdot>\<^sub>S ((\<mu> \<star> a) \<star> a)"
proof -
have "(\<mu> \<star> a \<star> a) \<cdot>\<^sub>S \<a>[S.dom \<mu>, a, a] = (\<mu> \<star> a \<star> a) \<cdot> \<a>[S.dom \<mu>, a, a]"
- using assms 3 S.ide_dom characteristic_iso(1) S.in_hom_char
+ using assms 3 S.ide_dom characteristic_iso(1) S.in_hom_char\<^sub>S\<^sub>b\<^sub>C
S.comp_char S.dom_simp
by fastforce
also have "... = \<a>[S.cod \<mu>, a, a] \<cdot> ((\<mu> \<star> a) \<star> a)"
using assms weak_unit_a assoc_naturality\<^sub>A\<^sub>W\<^sub>C [of \<mu> a a] S.dom_simp S.cod_simp
- weak_unit_self_composable S.arr_char right_def
+ weak_unit_self_composable S.arr_char\<^sub>S\<^sub>b\<^sub>C right_def
by simp
also have "... = \<a>[S.cod \<mu>, a, a] \<cdot>\<^sub>S ((\<mu> \<star> a) \<star> a)"
- using S.in_hom_char S.comp_char
- by (metis 2 4 5 6 R.preserves_arr S.seq_char)
+ using S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.comp_char
+ by (metis 2 4 5 6 R.preserves_arr S.seq_char\<^sub>S\<^sub>b\<^sub>C)
finally show ?thesis by blast
qed
thus ?thesis by argo
qed
also have "... = ((S.cod \<mu> \<star> \<iota>) \<cdot>\<^sub>S (\<mu> \<star> a \<star> a)) \<cdot>\<^sub>S \<a>[S.dom \<mu>, a, a]"
using S.comp_assoc by auto
also have "... = ((\<mu> \<star> a) \<cdot>\<^sub>S (S.dom \<mu> \<star> \<iota>)) \<cdot>\<^sub>S \<a>[S.dom \<mu>, a, a]"
proof -
have "\<mu> \<star> a \<star> a \<noteq> null"
using 3 S.not_arr_null by auto
moreover have "S.dom \<mu> \<star> \<iota> \<noteq> null"
using assms S.not_arr_null
- by (metis S.dom_char \<iota>_in_hom calculation hom_connected(1-2) in_homE)
+ by (metis S.dom_char\<^sub>S\<^sub>b\<^sub>C \<iota>_in_hom calculation hom_connected(1-2) in_homE)
ultimately have "(S.cod \<mu> \<star> \<iota>) \<cdot>\<^sub>S (\<mu> \<star> a \<star> a) = (\<mu> \<star> a) \<cdot>\<^sub>S (S.dom \<mu> \<star> \<iota>)"
using assms weak_unit_a iso_unit S.comp_arr_dom S.comp_cod_arr
interchange [of "S.cod \<mu>" \<mu> \<iota> "a \<star> a"] interchange [of \<mu> "S.dom \<mu>" a \<iota>]
by auto
thus ?thesis by argo
qed
also have "... = (\<mu> \<star> a) \<cdot>\<^sub>S (S.dom \<mu> \<star> \<iota>) \<cdot>\<^sub>S \<a>[S.dom \<mu>, a, a]"
using S.comp_assoc by auto
also have "... = R \<mu> \<cdot>\<^sub>S R \<r>[S.dom \<mu>]"
using assms runit_char(2) S.ide_dom H\<^sub>R_def by auto
also have "... = R (\<mu> \<cdot>\<^sub>S \<r>[S.dom \<mu>])"
using assms S.arr_iff_in_hom [of \<mu>] runit_char(1) S.ide_dom by fastforce
finally show ?thesis by blast
qed
ultimately show "\<r>[S.cod \<mu>] \<cdot>\<^sub>S (R \<mu>) = \<mu> \<cdot>\<^sub>S \<r>[S.dom \<mu>]"
using R.is_faithful by blast
qed
abbreviation \<rr>
where "\<rr> \<mu> \<equiv> if S.arr \<mu> then \<mu> \<cdot>\<^sub>S \<r>[S.dom \<mu>] else null"
interpretation \<rr>: natural_transformation S.comp S.comp R S.map \<rr>
proof -
interpret \<rr>: transformation_by_components S.comp S.comp R S.map runit
using runit_char(1) runit_naturality by unfold_locales simp_all
have "\<rr>.map = \<rr>"
using \<rr>.is_extensional \<rr>.map_def \<rr>.naturality \<rr>.map_simp_ide S.ide_dom S.ide_cod
S.map_def
by auto
thus "natural_transformation S.comp S.comp R S.map \<rr>"
using \<rr>.natural_transformation_axioms by auto
qed
lemma natural_transformation_\<rr>:
shows "natural_transformation S.comp S.comp R S.map \<rr>" ..
interpretation \<rr>: natural_isomorphism S.comp S.comp R S.map \<rr>
using S.ide_is_iso iso_runit runit_char(1) S.isos_compose
by unfold_locales force
lemma natural_isomorphism_\<rr>:
shows "natural_isomorphism S.comp S.comp R S.map \<rr>" ..
interpretation R: equivalence_functor S.comp S.comp R
using natural_isomorphism_\<rr> 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>[R f] = R \<r>[f]"
using assms runit_char(1) R.preserves_hom [of "\<r>[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 \<open>
Symmetric results hold for the subcategory of all arrows composable on the left with
a specified weak unit \<open>b\<close>. This yields the \emph{left unitors}.
\<close>
locale left_hom_with_unit =
associative_weak_composition V H \<a> +
left_hom V H b
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a comp" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<iota> :: 'a
and b :: 'a +
assumes weak_unit_b: "weak_unit b"
and \<iota>_in_hom: "\<guillemotleft>\<iota> : b \<star> b \<Rightarrow> b\<guillemotright>"
and iso_\<iota>: "iso \<iota>"
begin
abbreviation L
where "L \<equiv> 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 ("\<l>[_]")
where "lunit f \<equiv> THE \<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow>\<^sub>S f\<guillemotright> \<and> L \<mu> = (\<iota> \<star> f) \<cdot>\<^sub>S (inv \<a>[b, b, f])"
lemma iso_unit:
shows "S.iso \<iota>" and "\<guillemotleft>\<iota> : b \<star> b \<Rightarrow>\<^sub>S b\<guillemotright>"
proof -
show "\<guillemotleft>\<iota> : b \<star> b \<Rightarrow>\<^sub>S b\<guillemotright>"
- 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
+ using weak_unit_b S.ide_char S.arr_char\<^sub>S\<^sub>b\<^sub>C left_def weak_unit_self_composable
+ S.ideD(1) L.preserves_arr H\<^sub>L_def S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C left_def
\<iota>_in_hom S.ideD(1) hom_connected(4) in_homE
by metis
thus "S.iso \<iota>"
using iso_\<iota> iso_char by blast
qed
lemma characteristic_iso:
assumes "S.ide f"
shows "\<guillemotleft>inv \<a>[b, b, f] : b \<star> b \<star> f \<Rightarrow>\<^sub>S (b \<star> b) \<star> f\<guillemotright>"
and "\<guillemotleft>\<iota> \<star> f : (b \<star> b) \<star> f \<Rightarrow>\<^sub>S b \<star> f\<guillemotright>"
and "\<guillemotleft>(\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f] : L (L f) \<Rightarrow>\<^sub>S L f\<guillemotright>"
and "S.iso ((\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f])"
proof -
have f: "S.ide f \<and> ide f"
- using assms S.ide_char by simp
+ using assms S.ide_char\<^sub>S\<^sub>b\<^sub>C by simp
have b: "weak_unit b \<and> ide b \<and> S.ide b"
- using weak_unit_b S.ide_char weak_unit_def S.arr_char left_def
+ using weak_unit_b S.ide_char\<^sub>S\<^sub>b\<^sub>C weak_unit_def S.arr_char\<^sub>S\<^sub>b\<^sub>C left_def
weak_unit_self_composable
by metis
have bf: "b \<star> f \<noteq> null \<and> b \<star> b \<star> b \<star> f \<noteq> 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 \<star> b \<noteq> 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: "\<iota> \<star> f \<noteq> null"
- using assms S.ide_char left_def S.arr_char hom_connected(3) \<iota>_in_hom
+ using assms S.ide_char left_def S.arr_char\<^sub>S\<^sub>b\<^sub>C hom_connected(3) \<iota>_in_hom
by auto
show assoc_in_hom: "\<guillemotleft>inv \<a>[b, b, f] : b \<star> b \<star> f \<Rightarrow>\<^sub>S (b \<star> b) \<star> f\<guillemotright>"
using b f bf bb hom_connected(2) [of b "inv \<a>[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))
+ by (metis S.arrI\<^sub>S\<^sub>b\<^sub>C S.cod_closed S.in_hom_char\<^sub>S\<^sub>b\<^sub>C assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C(3) assoc'_simps\<^sub>A\<^sub>W\<^sub>C(2-3))
show 1: "\<guillemotleft>\<iota> \<star> f : (b \<star> b) \<star> f \<Rightarrow>\<^sub>S b \<star> f\<guillemotright>"
using b f bf right_hcomp_closed
by (simp add: ib_f ide_in_hom iso_unit(2))
have unit_part: "\<guillemotleft>\<iota> \<star> f : (b \<star> b) \<star> f \<Rightarrow>\<^sub>S b \<star> f\<guillemotright> \<and> S.iso (\<iota> \<star> f)"
proof -
have "S.iso (\<iota> \<star> f)"
- using b f bf ib_f 1 VoV.arr_char VxV.inv_simp
- inv_in_hom hom_connected(1) [of "inv \<iota>" f] VoV.arr_char VoV.iso_char
+ using b f bf ib_f 1 VoV.arr_char\<^sub>S\<^sub>b\<^sub>C VxV.inv_simp
+ inv_in_hom hom_connected(1) [of "inv \<iota>" f] VoV.arr_char\<^sub>S\<^sub>b\<^sub>C VoV.iso_char\<^sub>S\<^sub>b\<^sub>C
preserves_iso iso_char iso_\<iota> S.dom_simp S.cod_simp
by auto
thus ?thesis using 1 by blast
qed
show "S.iso ((\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f])"
using assms b f bf bb hom_connected(2) [of b "inv \<a>[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
+ iso_assoc\<^sub>A\<^sub>W\<^sub>C iso_char S.arr_char\<^sub>S\<^sub>b\<^sub>C unit_part assoc_in_hom isos_compose
S.isos_compose S.seqI' by auto
show "\<guillemotleft>(\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f] : L (L f) \<Rightarrow>\<^sub>S L f\<guillemotright>"
unfolding H\<^sub>L_def using unit_part assoc_in_hom by blast
qed
lemma lunit_char:
assumes "S.ide f"
shows "\<guillemotleft>\<l>[f] : L f \<Rightarrow>\<^sub>S f\<guillemotright>" and "L \<l>[f] = (\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow>\<^sub>S f\<guillemotright> \<and> L \<mu> = (\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f]"
proof -
let ?P = "\<lambda>\<mu>. \<guillemotleft>\<mu> : L f \<Rightarrow>\<^sub>S f\<guillemotright> \<and> L \<mu> = (\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f]"
show "\<exists>!\<mu>. ?P \<mu>"
proof -
have "\<exists>\<mu>. ?P \<mu>"
proof -
have 1: "S.ide f"
- using assms S.ide_char S.arr_char by simp
+ using assms S.ide_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "\<forall>\<mu> \<mu>'. ?P \<mu> \<and> ?P \<mu>' \<longrightarrow> \<mu> = \<mu>'"
using L.is_faithful S.in_homE by metis
ultimately show ?thesis by blast
qed
hence "?P (THE \<mu>. ?P \<mu>)"
using theI' [of ?P] by fastforce
hence 1: "?P \<l>[f]"
unfolding lunit_def by blast
show "\<guillemotleft>\<l>[f] : L f \<Rightarrow>\<^sub>S f\<guillemotright>" using 1 by fast
show "L \<l>[f] = (\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f]" using 1 by fast
qed
lemma iso_lunit:
assumes "S.ide f"
shows "S.iso \<l>[f]"
using assms characteristic_iso(4) lunit_char L.reflects_iso by metis
lemma lunit_eqI:
assumes "\<guillemotleft>f : a \<Rightarrow>\<^sub>S b\<guillemotright>" and "\<guillemotleft>\<mu> : L f \<Rightarrow>\<^sub>S f\<guillemotright>"
and "L \<mu> = ((\<iota> \<star> f) \<cdot>\<^sub>S inv \<a>[b, b, f])"
shows "\<mu> = \<l>[f]"
using assms S.ide_cod lunit_char [of f] by blast
lemma lunit_naturality:
assumes "S.arr \<mu>"
shows "\<l>[S.cod \<mu>] \<cdot>\<^sub>S L \<mu> = \<mu> \<cdot>\<^sub>S \<l>[S.dom \<mu>]"
proof -
have 1: "\<guillemotleft>\<l>[S.cod \<mu>] \<cdot>\<^sub>S L \<mu> : L (S.dom \<mu>) \<Rightarrow>\<^sub>S S.cod \<mu>\<guillemotright>"
using assms lunit_char(1) [of "S.cod \<mu>"] S.ide_cod by blast
have "S.par (\<l>[S.cod \<mu>] \<cdot>\<^sub>S L \<mu>) (\<mu> \<cdot>\<^sub>S \<l>[S.dom \<mu>])"
using assms 1 S.ide_dom lunit_char(1)
by (metis S.comp_in_homI' S.in_homE)
moreover have "L (\<l>[S.cod \<mu>] \<cdot>\<^sub>S L \<mu>) = L (\<mu> \<cdot>\<^sub>S \<l>[S.dom \<mu>])"
proof -
have 2: "\<guillemotleft>b \<star> b \<star> \<mu> : b \<star> b \<star> S.dom \<mu> \<Rightarrow>\<^sub>S b \<star> b \<star> S.cod \<mu>\<guillemotright>"
- using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom [of \<mu>] S.arr_char
+ using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom [of \<mu>] S.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
have 3: "\<guillemotleft>(b \<star> b) \<star> \<mu> : (b \<star> b) \<star> S.dom \<mu> \<Rightarrow>\<^sub>S (b \<star> b) \<star> S.cod \<mu>\<guillemotright>"
- 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
+ using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom S.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis match_3 weak_unit_in_vhom weak_unit_self_left S.in_hom_char\<^sub>S\<^sub>b\<^sub>C
S.not_arr_null S.null_char right_hcomp_closed)
have "L (\<l>[S.cod \<mu>] \<cdot>\<^sub>S L \<mu>) = L \<l>[S.cod \<mu>] \<cdot>\<^sub>S L (L \<mu>)"
using assms 1 L.as_nat_trans.preserves_comp_2 by blast
also have "... = ((\<iota> \<star> S.cod \<mu>) \<cdot>\<^sub>S inv \<a>[b, b, S.cod \<mu>]) \<cdot>\<^sub>S (b \<star> b \<star> \<mu>)"
using assms L.preserves_arr lunit_char S.ide_cod H\<^sub>L_def by auto
also have "... = (\<iota> \<star> S.cod \<mu>) \<cdot>\<^sub>S inv \<a>[b, b, S.cod \<mu>] \<cdot>\<^sub>S (b \<star> b \<star> \<mu>)"
using S.comp_assoc by auto
also have "... = (\<iota> \<star> S.cod \<mu>) \<cdot>\<^sub>S ((b \<star> b) \<star> \<mu>) \<cdot>\<^sub>S inv \<a>[b, b, S.dom \<mu>]"
proof -
have "inv \<a>[b, b, S.cod \<mu>] \<cdot>\<^sub>S (b \<star> b \<star> \<mu>) = ((b \<star> b) \<star> \<mu>) \<cdot>\<^sub>S inv \<a>[b, b, S.dom \<mu>]"
proof -
have "((b \<star> b) \<star> \<mu>) \<cdot>\<^sub>S inv \<a>[b, b, S.dom \<mu>] = ((b \<star> b) \<star> \<mu>) \<cdot> inv \<a>[b, b, S.dom \<mu>]"
- using assms 3 S.in_hom_char S.comp_char [of "(b \<star> b) \<star> \<mu>" "inv \<a>[b, b, S.dom \<mu>]"]
+ using assms 3 S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.comp_char [of "(b \<star> b) \<star> \<mu>" "inv \<a>[b, b, S.dom \<mu>]"]
by (metis S.ide_dom characteristic_iso(1) ext)
also have "... = inv \<a>[b, b, S.cod \<mu>] \<cdot> (b \<star> b \<star> \<mu>)"
using assms weak_unit_b assoc'_naturality\<^sub>A\<^sub>W\<^sub>C [of b b \<mu>] S.dom_simp S.cod_simp
- weak_unit_self_composable S.arr_char left_def
+ weak_unit_self_composable S.arr_char\<^sub>S\<^sub>b\<^sub>C left_def
by simp
also have "... = inv \<a>[b, b, S.cod \<mu>] \<cdot>\<^sub>S (b \<star> b \<star> \<mu>)"
- using assms 2 S.in_hom_char S.comp_char
+ using assms 2 S.in_hom_char\<^sub>S\<^sub>b\<^sub>C 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 "... = ((\<iota> \<star> S.cod \<mu>) \<cdot>\<^sub>S ((b \<star> b) \<star> \<mu>)) \<cdot>\<^sub>S inv \<a>[b, b, S.dom \<mu>]"
using S.comp_assoc by auto
also have "... = ((b \<star> \<mu>) \<cdot>\<^sub>S (\<iota> \<star> S.dom \<mu>)) \<cdot>\<^sub>S inv \<a>[b, b, S.dom \<mu>]"
proof -
have "(b \<star> b) \<star> \<mu> \<noteq> null"
using 3 S.not_arr_null by (elim S.in_homE, auto)
moreover have "\<iota> \<star> S.dom \<mu> \<noteq> null"
using assms S.not_arr_null
- by (metis S.dom_char \<iota>_in_hom calculation hom_connected(1-2) in_homE)
+ by (metis S.dom_char\<^sub>S\<^sub>b\<^sub>C \<iota>_in_hom calculation hom_connected(1-2) in_homE)
ultimately have "(\<iota> \<star> S.cod \<mu>) \<cdot>\<^sub>S ((b \<star> b) \<star> \<mu>) = (b \<star> \<mu>) \<cdot>\<^sub>S (\<iota> \<star> S.dom \<mu>)"
using assms weak_unit_b iso_unit S.comp_arr_dom S.comp_cod_arr
interchange [of \<iota> "b \<star> b" "S.cod \<mu>" \<mu> ] interchange [of b \<iota> \<mu> "S.dom \<mu>"]
by auto
thus ?thesis by argo
qed
also have "... = (b \<star> \<mu>) \<cdot>\<^sub>S (\<iota> \<star> S.dom \<mu>) \<cdot>\<^sub>S inv \<a>[b, b, S.dom \<mu>]"
using S.comp_assoc by auto
also have "... = L \<mu> \<cdot>\<^sub>S L \<l>[S.dom \<mu>]"
using assms lunit_char(2) S.ide_dom H\<^sub>L_def by auto
also have "... = L (\<mu> \<cdot>\<^sub>S \<l>[S.dom \<mu>])"
using assms S.arr_iff_in_hom [of \<mu>] lunit_char(1) S.ide_dom S.seqI
by fastforce
finally show ?thesis by blast
qed
ultimately show "\<l>[S.cod \<mu>] \<cdot>\<^sub>S L \<mu> = \<mu> \<cdot>\<^sub>S \<l>[S.dom \<mu>]"
using L.is_faithful by blast
qed
abbreviation \<ll>
where "\<ll> \<mu> \<equiv> if S.arr \<mu> then \<mu> \<cdot>\<^sub>S \<l>[S.dom \<mu>] else null"
interpretation \<ll>: natural_transformation S.comp S.comp L S.map \<ll>
proof -
interpret \<ll>: transformation_by_components S.comp S.comp L S.map lunit
using lunit_char(1) lunit_naturality by (unfold_locales, simp_all)
have "\<ll>.map = \<ll>"
using \<ll>.is_extensional \<ll>.map_def \<ll>.naturality \<ll>.map_simp_ide S.ide_dom S.ide_cod
S.map_def
by auto
thus "natural_transformation S.comp S.comp L S.map \<ll>"
using \<ll>.natural_transformation_axioms by auto
qed
lemma natural_transformation_\<ll>:
shows "natural_transformation S.comp S.comp L S.map \<ll>" ..
interpretation \<ll>: natural_isomorphism S.comp S.comp L S.map \<ll>
using S.ide_is_iso iso_lunit lunit_char(1) S.isos_compose
by (unfold_locales, force)
lemma natural_isomorphism_\<ll>:
shows "natural_isomorphism S.comp S.comp L S.map \<ll>" ..
interpretation L: equivalence_functor S.comp S.comp L
using natural_isomorphism_\<ll> 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>[L f] = L \<l>[f]"
using assms lunit_char(1) L.preserves_hom [of "\<l>[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 \<open>
A \emph{prebicategory} is an associative weak composition satisfying the additional assumption
that every arrow has a source and a target.
\<close>
locale prebicategory =
associative_weak_composition +
assumes arr_has_source: "arr \<mu> \<Longrightarrow> sources \<mu> \<noteq> {}"
and arr_has_target: "arr \<mu> \<Longrightarrow> targets \<mu> \<noteq> {}"
begin
lemma arr_iff_has_src:
shows "arr \<mu> \<longleftrightarrow> sources \<mu> \<noteq> {}"
using arr_has_source composable_implies_arr by auto
lemma arr_iff_has_trg:
shows "arr \<mu> \<longleftrightarrow> targets \<mu> \<noteq> {}"
using arr_has_target composable_implies_arr by auto
end
text \<open>
The horizontal composition of a prebicategory is regular.
\<close>
sublocale prebicategory \<subseteq> regular_weak_composition V H
proof
show "\<And>a f. a \<in> sources f \<Longrightarrow> ide f \<Longrightarrow> f \<star> a \<cong> f"
proof -
fix a f
assume a: "a \<in> sources f" and f: "ide f"
interpret Right_a: subcategory V \<open>right a\<close>
using a right_hom_is_subcategory weak_unit_self_composable by force
interpret Right_a: right_hom_with_unit V H \<a> \<open>some_unit a\<close> a
using a iso_some_unit by (unfold_locales, auto)
show "f \<star> a \<cong> f"
proof -
have "Right_a.ide f"
- using a f Right_a.ide_char Right_a.arr_char right_def by auto
+ using a f Right_a.ide_char\<^sub>S\<^sub>b\<^sub>C Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def by auto
hence "Right_a.iso (Right_a.runit f) \<and> (Right_a.runit f) \<in> Right_a.hom (f \<star> a) f"
using Right_a.iso_runit Right_a.runit_char(1) H\<^sub>R_def by simp
hence "iso (Right_a.runit f) \<and> (Right_a.runit f) \<in> hom (f \<star> a) f"
using Right_a.iso_char Right_a.hom_char by auto
thus ?thesis using f isomorphic_def by auto
qed
qed
show "\<And>b f. b \<in> targets f \<Longrightarrow> ide f \<Longrightarrow> b \<star> f \<cong> f"
proof -
fix b f
assume b: "b \<in> targets f" and f: "ide f"
interpret Left_b: subcategory V \<open>left b\<close>
using b left_hom_is_subcategory weak_unit_self_composable by force
interpret Left_b: left_hom_with_unit V H \<a> \<open>some_unit b\<close> b
using b iso_some_unit by (unfold_locales, auto)
show "b \<star> f \<cong> f"
proof -
have "Left_b.ide f"
- using b f Left_b.ide_char Left_b.arr_char left_def by auto
+ using b f Left_b.ide_char\<^sub>S\<^sub>b\<^sub>C Left_b.arr_char\<^sub>S\<^sub>b\<^sub>C left_def by auto
hence "Left_b.iso (Left_b.lunit f) \<and> (Left_b.lunit f) \<in> Left_b.hom (b \<star> 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) \<and> (Left_b.lunit f) \<in> hom (b \<star> f) f"
using Left_b.iso_char Left_b.hom_char by auto
thus ?thesis using isomorphic_def by auto
qed
qed
qed
text \<open>
The regularity allows us to show that, in a prebicategory, all sources of
a given arrow are isomorphic, and similarly for targets.
\<close>
context prebicategory
begin
lemma sources_are_isomorphic:
assumes "a \<in> sources \<mu>" and "a' \<in> sources \<mu>"
shows "a \<cong> a'"
proof -
have \<mu>: "arr \<mu>" using assms composable_implies_arr by auto
have "\<And>f. \<lbrakk> ide f; a \<in> sources f; a' \<in> sources f \<rbrakk> \<Longrightarrow> a \<cong> a'"
using \<mu> 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 \<mu>) \<and> a \<in> sources (dom \<mu>) \<and> a' \<in> sources (dom \<mu>)"
using assms \<mu> sources_dom by auto
ultimately show ?thesis by auto
qed
lemma targets_are_isomorphic:
assumes "b \<in> targets \<mu>" and "b' \<in> targets \<mu>"
shows "b \<cong> b'"
proof -
have \<mu>: "arr \<mu>" using assms composable_implies_arr by auto
have "\<And>f. \<lbrakk> ide f; b \<in> targets f; b' \<in> targets f \<rbrakk> \<Longrightarrow> b \<cong> b'"
by (metis connected_if_composable sources_are_isomorphic targetsD(3))
moreover have "ide (dom \<mu>) \<and> b \<in> targets (dom \<mu>) \<and> b' \<in> targets (dom \<mu>)"
using assms \<mu> targets_dom [of \<mu>] by auto
ultimately show ?thesis by auto
qed
text \<open>
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.
\<close>
interpretation H: partial_magma H
using is_partial_magma by auto
lemma isomorphism_respects_weak_units:
assumes "weak_unit a" and "a \<cong> a'"
shows "weak_unit a'"
proof -
obtain \<phi> where \<phi>: "iso \<phi> \<and> \<guillemotleft>\<phi> : a \<Rightarrow> a'\<guillemotright>"
using assms by auto
interpret Left_a: subcategory V \<open>left a\<close>
using assms left_hom_is_subcategory by fastforce
interpret Left_a: left_hom_with_unit V H \<a> \<open>some_unit a\<close> 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 \<a> \<open>some_unit a\<close> a
using assms iso_some_unit weak_unit_self_composable
apply unfold_locales by auto
have a': "ide a' \<and> a \<star> a' \<noteq> null \<and> a' \<star> a \<noteq> null \<and> a' \<star> a' \<noteq> null \<and>
\<phi> \<star> a' \<noteq> null \<and> Left_a.ide a'"
by (metis (no_types, lifting) Left_a.left_hom_axioms Right_a.weak_unit_a \<phi> assms(2)
ide_cod hom_connected(1) in_homE isomorphic_implies_equicomposable(1)
- left_def left_hom_def subcategory.ideI isomorphic_implies_equicomposable(2)
+ left_def left_hom_def subcategory.ideI\<^sub>S\<^sub>b\<^sub>C isomorphic_implies_equicomposable(2)
weak_unit_self_composable(3))
have iso: "a' \<star> a' \<cong> a'"
proof -
have 1: "Right a' = Right a"
using assms right_respects_isomorphic by simp
interpret Right_a': subcategory V \<open>right a'\<close>
using assms right_hom_is_subcategory by fastforce
(* TODO: The previous interpretation brings in unwanted notation for in_hom. *)
interpret Ra': endofunctor \<open>Right a'\<close> \<open>H\<^sub>R a'\<close>
using assms a' endofunctor_H\<^sub>R by auto
let ?\<psi> = "Left_a.lunit a' \<cdot> inv (\<phi> \<star> a')"
have "iso ?\<psi> \<and> \<guillemotleft>?\<psi> : a' \<star> a' \<Rightarrow> a'\<guillemotright>"
proof -
have "iso (Left_a.lunit a') \<and> \<guillemotleft>Left_a.lunit a' : a \<star> a' \<Rightarrow> a'\<guillemotright>"
using a' Left_a.lunit_char(1) Left_a.iso_lunit Left_a.iso_char
- Left_a.in_hom_char H\<^sub>L_def
+ Left_a.in_hom_char\<^sub>S\<^sub>b\<^sub>C H\<^sub>L_def
by auto
moreover have "iso (\<phi> \<star> a') \<and> \<guillemotleft>\<phi> \<star> a' : a \<star> a' \<Rightarrow> a' \<star> a'\<guillemotright>"
- using a' \<phi> 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
+ using a' \<phi> 1 Right_a'.arr_char\<^sub>S\<^sub>b\<^sub>C Right_a'.in_hom_char\<^sub>S\<^sub>b\<^sub>C Right_a.iso_char
+ right_def Ra'.preserves_iso Ra'.preserves_hom Right_a'.iso_char\<^sub>S\<^sub>b\<^sub>C
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 \<open>
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.
\<close>
have 1: "Left a' = Left a \<and> Right a' = Right a"
using assms left_respects_isomorphic right_respects_isomorphic by simp
interpret L: fully_faithful_functor \<open>Left a\<close> \<open>Left a\<close> \<open>H\<^sub>L a\<close>
using assms weak_unit_def by simp
interpret L': endofunctor \<open>Left a\<close> \<open>H\<^sub>L a'\<close>
using a' 1 endofunctor_H\<^sub>L [of a'] by auto
interpret \<Phi>: natural_isomorphism \<open>Left a\<close> \<open>Left a\<close> \<open>H\<^sub>L a\<close> \<open>H\<^sub>L a'\<close> \<open>H\<^sub>L \<phi>\<close>
proof
fix \<mu>
show "\<not> Left_a.arr \<mu> \<Longrightarrow> H\<^sub>L \<phi> \<mu> = Left_a.null"
using left_def \<phi> H\<^sub>L_def hom_connected(1) Left_a.null_char null_agreement
- Left_a.arr_char
+ Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
assume "Left_a.arr \<mu>"
hence \<mu>: "Left_a.arr \<mu> \<and> arr \<mu> \<and> a \<star> \<mu> \<noteq> null"
- using Left_a.arr_char left_def composable_implies_arr by simp
+ using Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def composable_implies_arr by simp
have 2: "\<phi> \<star> \<mu> \<noteq> null"
- using assms \<phi> \<mu> Left_a.arr_char left_def hom_connected by auto
+ using assms \<phi> \<mu> Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def hom_connected by auto
show "Left_a.dom (H\<^sub>L \<phi> \<mu>) = H\<^sub>L a (Left_a.dom \<mu>)"
- using assms 2 \<phi> \<mu> Left_a.arr_char left_def hom_connected(2) [of a \<phi>]
- weak_unit_self_composable match_4 Left_a.dom_char H\<^sub>L_def by auto
+ using assms 2 \<phi> \<mu> Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def hom_connected(2) [of a \<phi>]
+ weak_unit_self_composable match_4 Left_a.dom_char\<^sub>S\<^sub>b\<^sub>C H\<^sub>L_def by auto
show "Left_a.cod (H\<^sub>L \<phi> \<mu>) = H\<^sub>L a' (Left_a.cod \<mu>)"
- using assms 2 \<phi> \<mu> Left_a.arr_char left_def hom_connected(2) [of a \<phi>]
- weak_unit_self_composable match_4 Left_a.cod_char H\<^sub>L_def
+ using assms 2 \<phi> \<mu> Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def hom_connected(2) [of a \<phi>]
+ weak_unit_self_composable match_4 Left_a.cod_char\<^sub>S\<^sub>b\<^sub>C H\<^sub>L_def
by auto
show "Left_a.comp (H\<^sub>L a' \<mu>) (H\<^sub>L \<phi> (Left_a.dom \<mu>)) = H\<^sub>L \<phi> \<mu>"
proof -
have "Left_a.comp (H\<^sub>L a' \<mu>) (H\<^sub>L \<phi> (Left_a.dom \<mu>)) =
Left_a.comp (a' \<star> \<mu>) (\<phi> \<star> dom \<mu>)"
- using assms 1 2 \<phi> \<mu> Left_a.dom_char left_def H\<^sub>L_def by simp
+ using assms 1 2 \<phi> \<mu> Left_a.dom_char\<^sub>S\<^sub>b\<^sub>C left_def H\<^sub>L_def by simp
also have "... = (a' \<star> \<mu>) \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "Left_a.seq (a' \<star> \<mu>) (\<phi> \<star> dom \<mu>)"
proof (intro Left_a.seqI)
show 3: "Left_a.arr (\<phi> \<star> dom \<mu>)"
- using assms 2 \<phi> \<mu> Left_a.arr_char left_def
+ using assms 2 \<phi> \<mu> Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C 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' \<star> \<mu>)"
using \<mu> H\<^sub>L_def L'.preserves_arr by auto
show "Left_a.dom (a' \<star> \<mu>) = Left_a.cod (\<phi> \<star> dom \<mu>)"
- using a' \<phi> \<mu> 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)
+ using a' \<phi> \<mu> 2 3 4 Left_a.dom_char\<^sub>S\<^sub>b\<^sub>C Left_a.cod_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis Left_a.seqE Left_a.seq_char\<^sub>S\<^sub>b\<^sub>C 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
+ thus ?thesis using Left_a.comp_char Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def by auto
qed
also have "... = a' \<cdot> \<phi> \<star> \<mu> \<cdot> dom \<mu>"
using a' \<phi> \<mu> interchange hom_connected by auto
also have "... = \<phi> \<star> \<mu>"
using \<phi> \<mu> 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 \<phi> (Left_a.cod \<mu>)) (Left_a.L \<mu>) = H\<^sub>L \<phi> \<mu>"
proof -
have "Left_a.comp (H\<^sub>L \<phi> (Left_a.cod \<mu>)) (Left_a.L \<mu>) = Left_a.comp (\<phi> \<star> cod \<mu>) (a \<star> \<mu>)"
- using assms 1 2 \<phi> \<mu> Left_a.cod_char left_def H\<^sub>L_def by simp
+ using assms 1 2 \<phi> \<mu> Left_a.cod_char\<^sub>S\<^sub>b\<^sub>C left_def H\<^sub>L_def by simp
also have "... = (\<phi> \<star> cod \<mu>) \<cdot> (a \<star> \<mu>)"
proof -
have "Left_a.seq (\<phi> \<star> cod \<mu>) (a \<star> \<mu>)"
proof (intro Left_a.seqI)
show 3: "Left_a.arr (\<phi> \<star> cod \<mu>)"
- using \<phi> \<mu> 2 Left_a.arr_char left_def
+ using \<phi> \<mu> 2 Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<star> \<mu>)"
- using assms \<mu> Left_a.arr_char left_def
+ using assms \<mu> Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def
using H\<^sub>L_def L.preserves_arr by auto
show "Left_a.dom (\<phi> \<star> cod \<mu>) = Left_a.cod (a \<star> \<mu>)"
- using assms \<phi> \<mu> 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)
+ using assms \<phi> \<mu> 2 3 4 Left_a.dom_char\<^sub>S\<^sub>b\<^sub>C Left_a.cod_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis Left_a.seqE Left_a.seq_char\<^sub>S\<^sub>b\<^sub>C 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
+ thus ?thesis using Left_a.comp_char Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def by auto
qed
also have "... = \<phi> \<cdot> a \<star> cod \<mu> \<cdot> \<mu>"
using \<phi> \<mu> interchange hom_connected by auto
also have "... = \<phi> \<star> \<mu>"
using \<phi> \<mu> comp_arr_dom comp_cod_arr by auto
finally show ?thesis using H\<^sub>L_def by simp
qed
next
fix \<mu>
assume \<mu>: "Left_a.ide \<mu>"
have 1: "\<phi> \<star> \<mu> \<noteq> null"
- using assms \<phi> \<mu> Left_a.ide_char Left_a.arr_char left_def hom_connected by auto
+ using assms \<phi> \<mu> Left_a.ide_char Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def hom_connected by auto
show "Left_a.iso (H\<^sub>L \<phi> \<mu>)"
proof -
have "iso (\<phi> \<star> \<mu>)"
proof -
have "a \<in> sources \<phi> \<inter> targets \<mu>"
using assms \<phi> \<mu> 1 hom_connected weak_unit_self_composable
- Left_a.ide_char Left_a.arr_char left_def connected_if_composable
+ Left_a.ide_char Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_def connected_if_composable
by auto
thus ?thesis
- using \<phi> \<mu> Left_a.ide_char ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast
+ using \<phi> \<mu> Left_a.ide_char\<^sub>S\<^sub>b\<^sub>C ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast
qed
moreover have "left a (\<phi> \<star> \<mu>)"
using assms 1 \<phi> weak_unit_self_composable hom_connected(2) [of a \<phi>]
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
+ using Left_a.iso_char Left_a.arr_char\<^sub>S\<^sub>b\<^sub>C left_iff_left_inv Left_a.inv_char H\<^sub>L_def
by metis
qed
qed
interpret L': equivalence_functor \<open>Left a'\<close> \<open>Left a'\<close> \<open>H\<^sub>L a'\<close>
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))"
by (meson Left_a.natural_isomorphism_\<ll> \<Phi>.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 \<open>
Now we do the same for \<open>R'\<close>.
\<close>
interpret R: fully_faithful_functor \<open>Right a\<close> \<open>Right a\<close> \<open>H\<^sub>R a\<close>
using assms weak_unit_def by simp
interpret R': endofunctor \<open>Right a\<close> \<open>H\<^sub>R a'\<close>
using a' 1 endofunctor_H\<^sub>R [of a'] by auto
interpret \<Psi>: natural_isomorphism \<open>Right a\<close> \<open>Right a\<close> \<open>H\<^sub>R a\<close> \<open>H\<^sub>R a'\<close> \<open>H\<^sub>R \<phi>\<close>
proof
fix \<mu>
show "\<not> Right_a.arr \<mu> \<Longrightarrow> H\<^sub>R \<phi> \<mu> = Right_a.null"
- using right_def \<phi> H\<^sub>R_def hom_connected Right_a.null_char Right_a.arr_char
+ using right_def \<phi> H\<^sub>R_def hom_connected Right_a.null_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
assume "Right_a.arr \<mu>"
hence \<mu>: "Right_a.arr \<mu> \<and> arr \<mu> \<and> \<mu> \<star> a \<noteq> null"
- using Right_a.arr_char right_def composable_implies_arr by simp
+ using Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def composable_implies_arr by simp
have 2: "\<mu> \<star> \<phi> \<noteq> null"
- using assms \<phi> \<mu> Right_a.arr_char right_def hom_connected by auto
+ using assms \<phi> \<mu> Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def hom_connected by auto
show "Right_a.dom (H\<^sub>R \<phi> \<mu>) = H\<^sub>R a (Right_a.dom \<mu>)"
by (metis "2" H\<^sub>R_def R'.is_extensional Right_a.dom_simp Right_a.null_char
\<open>Right_a.arr \<mu>\<close> \<phi> a' hcomp_simps\<^sub>W\<^sub>C(2) in_homE match_3)
show "Right_a.cod (H\<^sub>R \<phi> \<mu>) = H\<^sub>R a' (Right_a.cod \<mu>)"
- using assms 2 a' \<phi> \<mu> Right_a.arr_char right_def hom_connected(3) [of \<phi> a]
- weak_unit_self_composable match_3 Right_a.cod_char H\<^sub>R_def
+ using assms 2 a' \<phi> \<mu> Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def hom_connected(3) [of \<phi> a]
+ weak_unit_self_composable match_3 Right_a.cod_char\<^sub>S\<^sub>b\<^sub>C H\<^sub>R_def
by auto
show "Right_a.comp (H\<^sub>R a' \<mu>) (H\<^sub>R \<phi> (Right_a.dom \<mu>)) = H\<^sub>R \<phi> \<mu>"
proof -
have "Right_a.comp (H\<^sub>R a' \<mu>) (H\<^sub>R \<phi> (Right_a.dom \<mu>)) =
Right_a.comp (\<mu> \<star> a') (dom \<mu> \<star> \<phi>)"
- using assms 1 2 \<phi> \<mu> Right_a.dom_char right_def H\<^sub>R_def by simp
+ using assms 1 2 \<phi> \<mu> Right_a.dom_char\<^sub>S\<^sub>b\<^sub>C right_def H\<^sub>R_def by simp
also have "... = (\<mu> \<star> a') \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "Right_a.seq (\<mu> \<star> a') (dom \<mu> \<star> \<phi>)"
proof (intro Right_a.seqI)
show 3: "Right_a.arr (dom \<mu> \<star> \<phi>)"
- using assms 2 \<phi> \<mu> Right_a.arr_char right_def
+ using assms 2 \<phi> \<mu> Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<mu> \<star> a')"
using \<mu> H\<^sub>R_def R'.preserves_arr by auto
show "Right_a.dom (\<mu> \<star> a') = Right_a.cod (dom \<mu> \<star> \<phi>)"
- using a' \<phi> \<mu> 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)
+ using a' \<phi> \<mu> 2 3 4 Right_a.dom_char\<^sub>S\<^sub>b\<^sub>C Right_a.cod_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis Right_a.seqE Right_a.seq_char\<^sub>S\<^sub>b\<^sub>C 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
+ thus ?thesis using Right_a.comp_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def by auto
qed
also have "... = \<mu> \<cdot> dom \<mu> \<star> a' \<cdot> \<phi>"
using a' \<phi> \<mu> interchange hom_connected by auto
also have "... = \<mu> \<star> \<phi>"
using \<phi> \<mu> 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 \<phi> (Right_a.cod \<mu>)) (Right_a.R \<mu>) = H\<^sub>R \<phi> \<mu>"
proof -
have "Right_a.comp (H\<^sub>R \<phi> (Right_a.cod \<mu>)) (Right_a.R \<mu>)
= Right_a.comp (cod \<mu> \<star> \<phi>) (\<mu> \<star> a)"
- using assms 1 2 \<phi> \<mu> Right_a.cod_char right_def H\<^sub>R_def by simp
+ using assms 1 2 \<phi> \<mu> Right_a.cod_char\<^sub>S\<^sub>b\<^sub>C right_def H\<^sub>R_def by simp
also have "... = (cod \<mu> \<star> \<phi>) \<cdot> (\<mu> \<star> a)"
proof -
have "Right_a.seq (cod \<mu> \<star> \<phi>) (\<mu> \<star> a)"
proof (intro Right_a.seqI)
show 3: "Right_a.arr (cod \<mu> \<star> \<phi>)"
- using \<phi> \<mu> 2 Right_a.arr_char right_def
+ using \<phi> \<mu> 2 Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<mu> \<star> a)"
- using assms \<mu> Right_a.arr_char right_def
+ using assms \<mu> Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def
using H\<^sub>R_def R.preserves_arr by auto
show "Right_a.dom (cod \<mu> \<star> \<phi>) = Right_a.cod (\<mu> \<star> a)"
- using assms \<phi> \<mu> 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)
+ using assms \<phi> \<mu> 2 3 4 Right_a.dom_char\<^sub>S\<^sub>b\<^sub>C Right_a.cod_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis Right_a.seqE Right_a.seq_char\<^sub>S\<^sub>b\<^sub>C 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
+ thus ?thesis using Right_a.comp_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def by auto
qed
also have "... = cod \<mu> \<cdot> \<mu> \<star> \<phi> \<cdot> a"
using \<phi> \<mu> interchange hom_connected by auto
also have "... = \<mu> \<star> \<phi>"
using \<phi> \<mu> comp_arr_dom comp_cod_arr by auto
finally show ?thesis using H\<^sub>R_def by simp
qed
next
fix \<mu>
assume \<mu>: "Right_a.ide \<mu>"
have 1: "\<mu> \<star> \<phi> \<noteq> null"
- using assms \<phi> \<mu> Right_a.ide_char Right_a.arr_char right_def hom_connected by auto
+ using assms \<phi> \<mu> Right_a.ide_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def hom_connected by auto
show "Right_a.iso (H\<^sub>R \<phi> \<mu>)"
proof -
have "iso (\<mu> \<star> \<phi>)"
proof -
have "a \<in> targets \<phi> \<inter> sources \<mu>"
using assms \<phi> \<mu> 1 hom_connected weak_unit_self_composable
- Right_a.ide_char Right_a.arr_char right_def connected_if_composable
+ Right_a.ide_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_def connected_if_composable
by (metis (full_types) IntI targetsI)
thus ?thesis
- using \<phi> \<mu> Right_a.ide_char ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast
+ using \<phi> \<mu> Right_a.ide_char\<^sub>S\<^sub>b\<^sub>C ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast
qed
moreover have "right a (\<mu> \<star> \<phi>)"
using assms 1 \<phi> weak_unit_self_composable hom_connected(1) [of \<phi> 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
+ using Right_a.iso_char Right_a.arr_char\<^sub>S\<^sub>b\<^sub>C right_iff_right_inv
Right_a.inv_char H\<^sub>R_def
by metis
qed
qed
interpret R': equivalence_functor \<open>Right a'\<close> \<open>Right a'\<close> \<open>H\<^sub>R a'\<close>
proof -
have "naturally_isomorphic (Right a') (Right a') (H\<^sub>R a')
(identity_functor.map (Right a'))"
proof -
have "naturally_isomorphic (Right a) (Right a) (H\<^sub>R a') Right_a.map"
by (meson Right_a.natural_isomorphism_\<rr> \<Psi>.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 \<in> sources \<mu>" and "a \<cong> a'"
shows "a' \<in> sources \<mu>"
using assms isomorphism_respects_weak_units isomorphic_implies_equicomposable
by blast
lemma targets_iso_closed:
assumes "a \<in> targets \<mu>" and "a \<cong> a'"
shows "a' \<in> targets \<mu>"
using assms isomorphism_respects_weak_units isomorphic_implies_equicomposable
by blast
lemma sources_eqI:
assumes "sources \<mu> \<inter> sources \<nu> \<noteq> {}"
shows "sources \<mu> = sources \<nu>"
using assms sources_iso_closed sources_are_isomorphic by blast
lemma targets_eqI:
assumes "targets \<mu> \<inter> targets \<nu> \<noteq> {}"
shows "targets \<mu> = targets \<nu>"
using assms targets_iso_closed targets_are_isomorphic by blast
text \<open>
The sets of sources and targets of a weak unit are isomorphism classes.
\<close>
lemma sources_char:
assumes "weak_unit a"
shows "sources a = {x. x \<cong> 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 \<cong> a}"
using assms targets_iso_closed weak_unit_iff_self_target targets_are_isomorphic
isomorphic_symmetric
by blast
end
section "Horizontal Homs"
text \<open>
Here we define a locale that axiomatizes a (vertical) category \<open>V\<close> that has been
punctuated into ``horizontal homs'' by the choice of idempotent endofunctors \<open>src\<close> and \<open>trg\<close>
that assign a specific ``source'' and ``target'' 1-cell to each of its arrows.
The functors \<open>src\<close> and \<open>trg\<close> are also subject to further conditions that constrain how
they commute with each other.
\<close>
locale horizontal_homs =
category V +
src: endofunctor V src +
trg: endofunctor V trg
for V :: "'a comp" (infixr "\<cdot>" 55)
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a" +
assumes ide_src [simp]: "arr \<mu> \<Longrightarrow> ide (src \<mu>)"
and ide_trg [simp]: "arr \<mu> \<Longrightarrow> ide (trg \<mu>)"
and src_src [simp]: "arr \<mu> \<Longrightarrow> src (src \<mu>) = src \<mu>"
and trg_trg [simp]: "arr \<mu> \<Longrightarrow> trg (trg \<mu>) = trg \<mu>"
and trg_src [simp]: "arr \<mu> \<Longrightarrow> trg (src \<mu>) = src \<mu>"
and src_trg [simp]: "arr \<mu> \<Longrightarrow> src (trg \<mu>) = trg \<mu>"
begin
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation in_hom ("\<guillemotleft>_ : _ \<Rightarrow> _\<guillemotright>")
text \<open>
We define an \emph{object} to be an arrow that is its own source
(or equivalently, its own target).
\<close>
definition obj
where "obj a \<equiv> arr a \<and> src a = a"
lemma obj_def':
shows "obj a \<longleftrightarrow> arr a \<and> 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 "\<lbrakk> ide a; src a = a; trg a = a \<rbrakk> \<Longrightarrow> 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 \<mu>"
shows "obj (src \<mu>)"
using assms objI_src by auto
lemma obj_trg [intro, simp]:
assumes "arr \<mu>"
shows "obj (trg \<mu>)"
using assms objI_trg by auto
definition in_hhom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
where "in_hhom \<mu> a b \<equiv> arr \<mu> \<and> src \<mu> = a \<and> trg \<mu> = b"
abbreviation hhom
where "hhom a b \<equiv> {\<mu>. \<guillemotleft>\<mu> : a \<rightarrow> b\<guillemotright>}"
abbreviation (input) hseq\<^sub>H\<^sub>H
where "hseq\<^sub>H\<^sub>H \<equiv> \<lambda>\<mu> \<nu>. arr \<mu> \<and> arr \<nu> \<and> src \<mu> = trg \<nu>"
lemma in_hhomI [intro, simp]:
assumes "arr \<mu>" and "src \<mu> = a" and "trg \<mu> = b"
shows "\<guillemotleft>\<mu> : a \<rightarrow> b\<guillemotright>"
using assms in_hhom_def by auto
lemma in_hhomE [elim]:
assumes "\<guillemotleft>\<mu> : a \<rightarrow> b\<guillemotright>"
and "\<lbrakk> arr \<mu>; obj a; obj b; src \<mu> = a; trg \<mu> = b \<rbrakk> \<Longrightarrow> 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 "\<guillemotleft>f : src f \<rightarrow> trg f\<guillemotright>" and "\<guillemotleft>f : f \<Rightarrow> f\<guillemotright>"
using assms by auto
lemma src_dom [simp]:
shows "src (dom \<mu>) = src \<mu>"
by (metis arr_dom_iff_arr obj_simps(4) obj_src src.is_extensional src.preserves_dom)
lemma src_cod [simp]:
shows "src (cod \<mu>) = src \<mu>"
by (metis arr_cod_iff_arr obj_simps(5) obj_src src.is_extensional src.preserves_cod)
lemma trg_dom [simp]:
shows "trg (dom \<mu>) = trg \<mu>"
by (metis arr_dom_iff_arr ide_char ide_trg trg.is_extensional trg.preserves_dom)
lemma trg_cod [simp]:
shows "trg (cod \<mu>) = trg \<mu>"
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 \<mu>) = src \<mu>"
by (metis dom_null ideD(2) ide_src src.is_extensional)
lemma cod_src [simp]:
shows "cod (src \<mu>) = src \<mu>"
by (metis cod_null ideD(3) ide_src src.is_extensional)
lemma dom_trg [simp]:
shows "dom (trg \<mu>) = trg \<mu>"
by (metis dom_null ideD(2) ide_trg trg.is_extensional)
lemma cod_trg [simp]:
shows "cod (trg \<mu>) = trg \<mu>"
by (metis cod_null ideD(3) ide_trg trg.is_extensional)
lemma vcomp_in_hhom [intro, simp]:
assumes "seq \<nu> \<mu>" and "src \<nu> = a" and "trg \<nu> = b"
shows "\<guillemotleft>\<nu> \<cdot> \<mu> : a \<rightarrow> b\<guillemotright>"
using assms src_cod [of "\<nu> \<cdot> \<mu>"] trg_cod [of "\<nu> \<cdot> \<mu>"] by auto
lemma src_vcomp [simp]:
assumes "seq \<nu> \<mu>"
shows "src (\<nu> \<cdot> \<mu>) = src \<nu>"
using assms src_cod [of "\<nu> \<cdot> \<mu>"] by auto
lemma trg_vcomp [simp]:
assumes "seq \<nu> \<mu>"
shows "trg (\<nu> \<cdot> \<mu>) = trg \<nu>"
using assms trg_cod [of "\<nu> \<cdot> \<mu>"] by auto
lemma vseq_implies_hpar:
assumes "seq \<nu> \<mu>"
shows "src \<nu> = src \<mu>" and "trg \<nu> = trg \<mu>"
using assms src_dom [of "\<nu> \<cdot> \<mu>"] trg_dom [of "\<nu> \<cdot> \<mu>"] src_cod [of "\<nu> \<cdot> \<mu>"]
trg_cod [of "\<nu> \<cdot> \<mu>"]
by auto
lemma vconn_implies_hpar:
assumes "\<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright>"
shows "src \<mu> = src f" and "trg \<mu> = trg f" and "src g = src f" and "trg g = trg f"
using assms by auto
lemma src_inv [simp]:
assumes "iso \<mu>"
shows "src (inv \<mu>) = src \<mu>"
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 \<mu>"
shows "trg (inv \<mu>) = trg \<mu>"
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 \<mu>" and "src \<mu> = a" and "trg \<mu> = b"
shows "\<guillemotleft>inv \<mu> : a \<rightarrow> b\<guillemotright>"
using assms iso_is_arr by simp
lemma hhom_is_subcategory:
shows "subcategory V (\<lambda>\<mu>. \<guillemotleft>\<mu> : a \<rightarrow> b\<guillemotright>)"
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 \<cong> b"
shows "a = b"
using assms isomorphic_def
by (metis dom_inv in_homE objE src_dom src_inv)
text \<open>
Having the functors \<open>src\<close> and \<open>trg\<close> allows us to form categories VV and VVV
of formally horizontally composable pairs and triples of arrows.
\<close>
sublocale VxV: product_category V V ..
sublocale VV: subcategory VxV.comp \<open>\<lambda>\<mu>\<nu>. hseq\<^sub>H\<^sub>H (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
by (unfold_locales, auto)
lemma subcategory_VV:
shows "subcategory VxV.comp (\<lambda>\<mu>\<nu>. hseq\<^sub>H\<^sub>H (fst \<mu>\<nu>) (snd \<mu>\<nu>))"
..
sublocale VxVxV: product_category V VxV.comp ..
sublocale VVV: subcategory VxVxV.comp
\<open>\<lambda>\<tau>\<mu>\<nu>. arr (fst \<tau>\<mu>\<nu>) \<and> VV.arr (snd \<tau>\<mu>\<nu>) \<and>
src (fst \<tau>\<mu>\<nu>) = trg (fst (snd \<tau>\<mu>\<nu>))\<close>
- using VV.arr_char
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (unfold_locales, auto)
lemma subcategory_VVV:
shows "subcategory VxVxV.comp
(\<lambda>\<tau>\<mu>\<nu>. arr (fst \<tau>\<mu>\<nu>) \<and> VV.arr (snd \<tau>\<mu>\<nu>) \<and>
src (fst \<tau>\<mu>\<nu>) = trg (fst (snd \<tau>\<mu>\<nu>)))"
..
end
subsection "Prebicategories with Homs"
text \<open>
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 \<open>\<mu>\<close> in fact lie in the set of sources and targets,
respectively, of \<open>\<mu>\<close>, such that horizontal composition respects the
chosen sources and targets, and such that if 2-cells \<open>\<mu>\<close> and \<open>\<nu>\<close> are
horizontally composable, then the chosen target of \<open>\<mu>\<close> coincides with
the chosen source of \<open>\<nu>\<close>.
\<close>
locale weak_composition_with_homs =
weak_composition +
horizontal_homs +
assumes src_in_sources: "arr \<mu> \<Longrightarrow> src \<mu> \<in> sources \<mu>"
and trg_in_targets: "arr \<mu> \<Longrightarrow> trg \<mu> \<in> targets \<mu>"
and src_hcomp': "\<nu> \<star> \<mu> \<noteq> null \<Longrightarrow> src (\<nu> \<star> \<mu>) = src \<mu>"
and trg_hcomp': "\<nu> \<star> \<mu> \<noteq> null \<Longrightarrow> trg (\<nu> \<star> \<mu>) = trg \<nu>"
and seq_if_composable: "\<nu> \<star> \<mu> \<noteq> null \<Longrightarrow> src \<nu> = trg \<mu>"
locale prebicategory_with_homs =
prebicategory +
weak_composition_with_homs
begin
lemma composable_char\<^sub>P\<^sub>B\<^sub>H:
shows "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> arr \<mu> \<and> arr \<nu> \<and> src \<nu> = trg \<mu>"
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 "\<guillemotleft>\<mu> : a \<rightarrow>\<^sub>W\<^sub>C b\<guillemotright>" and "\<guillemotleft>\<nu> : b \<rightarrow>\<^sub>W\<^sub>C c\<guillemotright>"
shows "\<guillemotleft>\<nu> \<star> \<mu> : a \<rightarrow>\<^sub>W\<^sub>C c\<guillemotright>"
and "\<guillemotleft>\<nu> \<star> \<mu> : dom \<nu> \<star> dom \<mu> \<Rightarrow> cod \<nu> \<star> cod \<mu>\<guillemotright>"
proof -
show "\<guillemotleft>\<nu> \<star> \<mu> : a \<rightarrow>\<^sub>W\<^sub>C c\<guillemotright>"
using assms sources_determine_composability sources_hcomp targets_hcomp by auto
thus "\<guillemotleft>\<nu> \<star> \<mu> : dom \<nu> \<star> dom \<mu> \<Rightarrow> cod \<nu> \<star> cod \<mu>\<guillemotright>"
using assms by auto
qed
text \<open>
In a prebicategory with homs, if \<open>a\<close> is an object (i.e. \<open>src a = a\<close> and \<open>trg a = a\<close>),
then \<open>a\<close> is a weak unit. The converse need not hold: there can be weak units that the
\<open>src\<close> and \<open>trg\<close> mappings send to other 1-cells in the same isomorphism class.
\<close>
lemma obj_is_weak_unit:
assumes "obj a"
shows "weak_unit a"
proof -
have "a \<in> sources a"
using assms objE src_in_sources ideD(1) by metis
thus ?thesis by auto
qed
end
subsection "Choosing Homs"
text \<open>
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''.
\<close>
context prebicategory
begin
definition rep
where "rep f \<equiv> SOME f'. f' \<in> { f'. f \<cong> f' }"
definition some_src
where "some_src \<mu> \<equiv> if arr \<mu> then rep (SOME a. a \<in> sources \<mu>) else null"
definition some_trg
where "some_trg \<mu> \<equiv> if arr \<mu> then rep (SOME b. b \<in> targets \<mu>) else null"
lemma isomorphic_ide_rep:
assumes "ide f"
shows "f \<cong> rep f"
proof -
have "\<exists>f'. f' \<in> { f'. f \<cong> 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 \<in> { f'. f \<cong> f' }"
using assms isomorphic_ide_rep by simp
have "{ f'. f \<cong> f' } = { f'. rep f \<cong> f' }"
proof -
have "\<And>f'. f \<cong> f' \<longleftrightarrow> rep f \<cong> f'"
proof
fix f'
assume f': "f \<cong> f'"
show "rep f \<cong> f'"
proof -
obtain \<phi> where \<phi>: "\<phi> \<in> hom f f' \<and> iso \<phi>"
using f' by auto
obtain \<psi> where \<psi>: "\<psi> \<in> hom f (rep f) \<and> iso \<psi>"
using assms isomorphic_ide_rep by blast
have "inv \<psi> \<in> hom (rep f) f \<and> iso (inv \<psi>)"
using \<psi> by simp
hence "iso (V \<phi> (inv \<psi>)) \<and> V \<phi> (inv \<psi>) \<in> hom (rep f) f'"
using \<phi> isos_compose by auto
thus ?thesis using isomorphic_def by auto
qed
next
fix f'
assume f': "rep f \<cong> f'"
show "f \<cong> 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' \<in> { f'. f \<cong> 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 \<mu>"
shows "some_src \<mu> \<in> sources \<mu>"
proof -
have 1: "(SOME a. a \<in> sources \<mu>) \<in> sources \<mu>"
using assms arr_iff_has_src someI_ex [of "\<lambda>a. a \<in> sources \<mu>"] by blast
moreover have "ide (SOME a. a \<in> sources \<mu>)"
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 \<mu>"
shows "some_trg \<mu> \<in> targets \<mu>"
proof -
have 1: "(SOME a. a \<in> targets \<mu>) \<in> targets \<mu>"
using assms arr_iff_has_trg someI_ex [of "\<lambda>a. a \<in> targets \<mu>"] by blast
moreover have "ide (SOME a. a \<in> targets \<mu>)"
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 \<mu>"
shows "some_src (dom \<mu>) = some_src \<mu>"
using assms some_src_def sources_dom by simp
lemma some_src_cod:
assumes "arr \<mu>"
shows "some_src (cod \<mu>) = some_src \<mu>"
using assms some_src_def sources_cod by simp
lemma some_trg_dom:
assumes "arr \<mu>"
shows "some_trg (dom \<mu>) = some_trg \<mu>"
using assms some_trg_def targets_dom by simp
lemma some_trg_cod:
assumes "arr \<mu>"
shows "some_trg (cod \<mu>) = some_trg \<mu>"
using assms some_trg_def targets_cod by simp
lemma ide_some_src:
assumes "arr \<mu>"
shows "ide (some_src \<mu>)"
using assms some_src_in_sources weak_unit_self_composable by blast
lemma ide_some_trg:
assumes "arr \<mu>"
shows "ide (some_trg \<mu>)"
using assms some_trg_in_targets weak_unit_self_composable by blast
lemma some_src_composable:
assumes "arr \<tau>"
shows "\<tau> \<star> \<mu> \<noteq> null \<longleftrightarrow> some_src \<tau> \<star> \<mu> \<noteq> null"
using assms some_src_in_sources sources_determine_composability by blast
lemma some_trg_composable:
assumes "arr \<sigma>"
shows "\<mu> \<star> \<sigma> \<noteq> null \<longleftrightarrow> \<mu> \<star> some_trg \<sigma> \<noteq> null"
using assms some_trg_in_targets targets_determine_composability by blast
lemma sources_some_src:
assumes "arr \<mu>"
shows "sources (some_src \<mu>) = sources \<mu>"
using assms sources_determine_composability some_src_in_sources by blast
lemma targets_some_trg:
assumes "arr \<mu>"
shows "targets (some_trg \<mu>) = targets \<mu>"
using assms targets_determine_composability some_trg_in_targets by blast
lemma src_some_src:
assumes "arr \<mu>"
shows "some_src (some_src \<mu>) = some_src \<mu>"
using assms some_src_def ide_some_src sources_some_src by force
lemma trg_some_trg:
assumes "arr \<mu>"
shows "some_trg (some_trg \<mu>) = some_trg \<mu>"
using assms some_trg_def ide_some_trg targets_some_trg by force
lemma sources_char':
assumes "arr \<mu>"
shows "a \<in> sources \<mu> \<longleftrightarrow> some_src \<mu> \<cong> a"
using assms some_src_in_sources sources_iso_closed sources_are_isomorphic by meson
lemma targets_char':
assumes "arr \<mu>"
shows "a \<in> targets \<mu> \<longleftrightarrow> some_trg \<mu> \<cong> a"
using assms some_trg_in_targets targets_iso_closed targets_are_isomorphic by meson
text \<open>
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.
\<close>
lemma composable_char\<^sub>P\<^sub>B:
shows "\<tau> \<star> \<sigma> \<noteq> null \<longleftrightarrow> arr \<sigma> \<and> arr \<tau> \<and> some_src \<tau> = some_trg \<sigma>"
proof
assume \<sigma>\<tau>: "\<tau> \<star> \<sigma> \<noteq> null"
show "arr \<sigma> \<and> arr \<tau> \<and> some_src \<tau> = some_trg \<sigma>"
using \<sigma>\<tau> composable_implies_arr connected_if_composable some_src_def some_trg_def
by force
next
assume \<sigma>\<tau>: "arr \<sigma> \<and> arr \<tau> \<and> some_src \<tau> = some_trg \<sigma>"
show "\<tau> \<star> \<sigma> \<noteq> null"
using \<sigma>\<tau> some_src_in_sources some_trg_composable by force
qed
text \<open>
A 1-cell is its own source if and only if it is its own target.
\<close>
lemma self_src_iff_self_trg:
assumes "ide a"
shows "a = some_src a \<longleftrightarrow> a = some_trg a"
proof
assume a: "a = some_src a"
have "weak_unit a \<and> a \<star> a \<noteq> 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 \<and> a \<star> a \<noteq> 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 \<mu>"
shows "some_trg (some_src \<mu>) = some_src \<mu>"
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 \<mu>"
shows "some_src (some_trg \<mu>) = some_trg \<mu>"
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 \<in> sources \<mu>" and "some_src a = a"
shows "some_src \<mu> = 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 \<in> targets \<mu>" and "some_trg b = b"
shows "some_trg \<mu> = 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 "\<tau> \<star> \<sigma> \<noteq> null"
shows "some_src (\<tau> \<star> \<sigma>) = some_src \<sigma>"
proof (intro some_src_eqI [of "some_src \<sigma>" "\<tau> \<star> \<sigma>"])
show "some_src (some_src \<sigma>) = some_src \<sigma>"
using assms src_some_src composable_implies_arr by simp
show "some_src \<sigma> \<in> sources (H \<tau> \<sigma>)"
using assms some_src_in_sources composable_char\<^sub>P\<^sub>B
by (simp add: sources_hcomp)
qed
lemma some_trg_comp:
assumes "\<tau> \<star> \<sigma> \<noteq> null"
shows "some_trg (\<tau> \<star> \<sigma>) = some_trg \<tau>"
proof (intro some_trg_eqI [of "some_trg \<tau>" "\<tau> \<star> \<sigma>"])
show "some_trg (some_trg \<tau>) = some_trg \<tau>"
using assms trg_some_trg composable_implies_arr by simp
show "some_trg \<tau> \<in> targets (H \<tau> \<sigma>)"
using assms some_trg_in_targets composable_char\<^sub>P\<^sub>B
by (simp add: targets_hcomp)
qed
text \<open>
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.
\<close>
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 \<nu> \<mu>
assume \<mu>\<nu>: "seq \<nu> \<mu>"
show "some_src (\<nu> \<cdot> \<mu>) = some_src \<nu> \<cdot> some_src \<mu>"
using \<mu>\<nu> some_src_dom [of "\<nu> \<cdot> \<mu>"] some_src_dom some_src_cod [of "\<nu> \<cdot> \<mu>"]
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 \<nu> \<mu>
assume \<mu>\<nu>: "seq \<nu> \<mu>"
show "some_trg (\<nu> \<cdot> \<mu>) = some_trg \<nu> \<cdot> some_trg \<mu>"
using \<mu>\<nu> some_trg_dom [of "\<nu> \<cdot> \<mu>"] some_trg_dom some_trg_cod [of "\<nu> \<cdot> \<mu>"]
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 \<a> some_src some_trg"
..
end
subsection "Choosing Units"
text \<open>
A \emph{prebicategory with units} is a prebicategory equipped with a choice,
for each weak unit \<open>a\<close>, of a ``unit isomorphism'' \<open>\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>\<close>.
\<close>
locale prebicategory_with_units =
prebicategory V H \<a> +
weak_composition V H
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a comp" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]") +
assumes unit_in_vhom\<^sub>P\<^sub>B\<^sub>U: "weak_unit a \<Longrightarrow> \<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
and iso_unit\<^sub>P\<^sub>B\<^sub>U: "weak_unit a \<Longrightarrow> iso \<i>[a]"
begin
lemma unit_in_hom\<^sub>P\<^sub>B\<^sub>U:
assumes "weak_unit a"
shows "\<guillemotleft>\<i>[a] : a \<rightarrow>\<^sub>W\<^sub>C a\<guillemotright>" and "\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
using assms unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by auto
show "\<guillemotleft>\<i>[a] : a \<rightarrow>\<^sub>W\<^sub>C a\<guillemotright>"
using assms 1 weak_unit_iff_self_source weak_unit_iff_self_target
sources_cod [of "\<i>[a]"] targets_cod [of "\<i>[a]"]
by (elim in_homE, auto)
qed
lemma unit_simps [simp]:
assumes "weak_unit a"
shows "arr \<i>[a]" and "dom \<i>[a] = a \<star> a" and "cod \<i>[a] = a"
using assms unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by auto
end
text \<open>
Every prebicategory extends to a prebicategory with units, simply by choosing the
unit isomorphisms arbitrarily.
\<close>
context prebicategory
begin
proposition extends_to_prebicategory_with_units:
shows "prebicategory_with_units V H \<a> some_unit"
using iso_some_unit by (unfold_locales, auto)
end
subsection "Horizontal Composition"
text \<open>
The following locale axiomatizes a (vertical) category \<open>V\<close> with horizontal homs,
which in addition has been equipped with a functorial operation \<open>H\<close> of
horizontal composition from \<open>VV\<close> to \<open>V\<close>, assumed to preserve source and target.
\<close>
locale horizontal_composition =
horizontal_homs V src trg +
H: "functor" VV.comp V \<open>\<lambda>\<mu>\<nu>. H (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a" +
assumes src_hcomp: "arr (\<mu> \<star> \<nu>) \<Longrightarrow> src (\<mu> \<star> \<nu>) = src \<nu>"
and trg_hcomp: "arr (\<mu> \<star> \<nu>) \<Longrightarrow> trg (\<mu> \<star> \<nu>) = trg \<mu>"
begin
(* TODO: Why does this get re-introduced? *)
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text \<open>
\<open>H\<close> is a partial magma, which shares its null with \<open>V\<close>.
\<close>
lemma is_partial_magma:
shows "partial_magma H" and "partial_magma.null H = null"
proof -
have 1: "\<forall>f. null \<star> f = null \<and> f \<star> null = null"
- using H.is_extensional VV.arr_char not_arr_null by auto
+ using H.is_extensional VV.arr_char\<^sub>S\<^sub>b\<^sub>C not_arr_null by auto
interpret H: partial_magma H
- using 1 VV.arr_char H.is_extensional not_arr_null
+ using 1 VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "\<lambda>n. \<forall>f. n \<star> f = n \<and> f \<star> n = n"]
by metis
qed
text \<open>
\textbf{Note:} The following is ``almost'' \<open>H.seq\<close>, but for that we would need
\<open>H.arr = V.arr\<close>.
This would be unreasonable to expect, in general, as the definition of \<open>H.arr\<close> is based
on ``strict'' units rather than weak units.
Later we will show that we do have \<open>H.arr = V.arr\<close> if the vertical category is discrete.
\<close>
abbreviation hseq
where "hseq \<nu> \<mu> \<equiv> arr (\<nu> \<star> \<mu>)"
lemma hseq_char:
shows "hseq \<nu> \<mu> \<longleftrightarrow> arr \<mu> \<and> arr \<nu> \<and> src \<nu> = trg \<mu>"
proof -
have "hseq \<nu> \<mu> \<longleftrightarrow> VV.arr (\<nu>, \<mu>)"
using H.is_extensional H.preserves_arr by force
also have "... \<longleftrightarrow> arr \<mu> \<and> arr \<nu> \<and> src \<nu> = trg \<mu>"
- using VV.arr_char by force
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C by force
finally show ?thesis by blast
qed
lemma hseq_char':
shows "hseq \<nu> \<mu> \<longleftrightarrow> \<nu> \<star> \<mu> \<noteq> null"
- using VV.arr_char H.preserves_arr H.is_extensional hseq_char [of \<nu> \<mu>] by auto
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.preserves_arr H.is_extensional hseq_char [of \<nu> \<mu>] by auto
lemma hseqI' [intro, simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "hseq \<nu> \<mu>"
using assms hseq_char by simp
lemma hseqI:
assumes "\<guillemotleft>\<mu> : a \<rightarrow> b\<guillemotright>" and "\<guillemotleft>\<nu> : b \<rightarrow> c\<guillemotright>"
shows "hseq \<nu> \<mu>"
using assms hseq_char by auto
lemma hseqE [elim]:
assumes "hseq \<nu> \<mu>"
and "arr \<mu> \<Longrightarrow> arr \<nu> \<Longrightarrow> src \<nu> = trg \<mu> \<Longrightarrow> T"
shows "T"
using assms hseq_char by simp
lemma hcomp_simps [simp]:
assumes "hseq \<nu> \<mu>"
shows "src (\<nu> \<star> \<mu>) = src \<mu>" and "trg (\<nu> \<star> \<mu>) = trg \<nu>"
and "dom (\<nu> \<star> \<mu>) = dom \<nu> \<star> dom \<mu>" and "cod (\<nu> \<star> \<mu>) = cod \<nu> \<star> cod \<mu>"
- 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
+ using assms VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_hcomp apply blast
+ using assms VV.arr_char\<^sub>S\<^sub>b\<^sub>C trg_hcomp apply blast
+ using assms VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.preserves_dom VV.dom_simp apply force
+ using assms VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.preserves_cod VV.cod_simp by force
lemma ide_hcomp [intro, simp]:
assumes "ide \<nu>" and "ide \<mu>" and "src \<nu> = trg \<mu>"
shows "ide (\<nu> \<star> \<mu>)"
- using assms VV.ide_char VV.arr_char H.preserves_ide [of "(\<nu>, \<mu>)"] by auto
+ using assms VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.preserves_ide [of "(\<nu>, \<mu>)"] by auto
lemma hcomp_in_hhom [intro]:
assumes "\<guillemotleft>\<mu> : a \<rightarrow> b\<guillemotright>" and "\<guillemotleft>\<nu> : b \<rightarrow> c\<guillemotright>"
shows "\<guillemotleft>\<nu> \<star> \<mu> : a \<rightarrow> c\<guillemotright>"
using assms hseq_char by fastforce
lemma hcomp_in_hhom' (* [simp] *):
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = a" and "trg \<nu> = c" and "src \<nu> = trg \<mu>"
shows "\<guillemotleft>\<nu> \<star> \<mu> : a \<rightarrow> c\<guillemotright>"
using assms hseq_char by fastforce
lemma hcomp_in_hhomE [elim]:
assumes "\<guillemotleft>\<nu> \<star> \<mu> : a \<rightarrow> c\<guillemotright>"
and "\<lbrakk> arr \<mu>; arr \<nu>; src \<nu> = trg \<mu>; src \<mu> = a; trg \<nu> = c \<rbrakk> \<Longrightarrow> T"
shows T
using assms in_hhom_def by fastforce
lemma hcomp_in_vhom [intro]:
assumes "\<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright>" and "\<guillemotleft>\<nu> : h \<Rightarrow> k\<guillemotright>" and "src h = trg f"
shows "\<guillemotleft>\<nu> \<star> \<mu> : h \<star> f \<Rightarrow> k \<star> g\<guillemotright>"
using assms by fastforce
lemma hcomp_in_vhom' (* [simp] *):
assumes "hseq \<nu> \<mu>"
and "dom \<mu> = f" and "dom \<nu> = h" and "cod \<mu> = g" and "cod \<nu> = k"
assumes "\<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright>" and "\<guillemotleft>\<nu> : h \<Rightarrow> k\<guillemotright>" and "src h = trg f"
shows "\<guillemotleft>\<nu> \<star> \<mu> : h \<star> f \<Rightarrow> k \<star> g\<guillemotright>"
using assms by fastforce
lemma hcomp_in_vhomE [elim]:
assumes "\<guillemotleft>\<nu> \<star> \<mu> : f \<Rightarrow> g\<guillemotright>"
and "\<lbrakk> arr \<mu>; arr \<nu>; src \<nu> = trg \<mu>; src \<mu> = src f; src \<mu> = src g;
trg \<nu> = trg f; trg \<nu> = trg g \<rbrakk> \<Longrightarrow> 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 \<open>
A horizontal composition yields a weak composition by simply forgetting
the \<open>src\<close> and \<open>trg\<close> functors.
\<close>
lemma match_1:
assumes "\<nu> \<star> \<mu> \<noteq> null" and "(\<nu> \<star> \<mu>) \<star> \<tau> \<noteq> null"
shows "\<mu> \<star> \<tau> \<noteq> null"
- using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto
+ using assms H.is_extensional not_arr_null VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char hseq_char' by auto
lemma match_2:
assumes "\<nu> \<star> (\<mu> \<star> \<tau>) \<noteq> null" and "\<mu> \<star> \<tau> \<noteq> null"
shows "\<nu> \<star> \<mu> \<noteq> null"
- using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto
+ using assms H.is_extensional not_arr_null VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char hseq_char' by auto
lemma match_3:
assumes "\<mu> \<star> \<tau> \<noteq> null" and "\<nu> \<star> \<mu> \<noteq> null"
shows "(\<nu> \<star> \<mu>) \<star> \<tau> \<noteq> null"
- using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto
+ using assms H.is_extensional not_arr_null VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char hseq_char' by auto
lemma match_4:
assumes "\<mu> \<star> \<tau> \<noteq> null" and "\<nu> \<star> \<mu> \<noteq> null"
shows "\<nu> \<star> (\<mu> \<star> \<tau>) \<noteq> null"
- using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto
+ using assms H.is_extensional not_arr_null VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char hseq_char' by auto
lemma left_connected:
assumes "seq \<nu> \<nu>'"
shows "\<nu> \<star> \<mu> \<noteq> null \<longleftrightarrow> \<nu>' \<star> \<mu> \<noteq> null"
- using assms H.is_extensional not_arr_null VV.arr_char hseq_char'
+ using assms H.is_extensional not_arr_null VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char'
by (metis hseq_char seqE vseq_implies_hpar(1))
lemma right_connected:
assumes "seq \<mu> \<mu>'"
shows "H \<nu> \<mu> \<noteq> null \<longleftrightarrow> H \<nu> \<mu>' \<noteq> null"
- using assms H.is_extensional not_arr_null VV.arr_char hseq_char'
+ using assms H.is_extensional not_arr_null VV.arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char'
by (metis hseq_char seqE vseq_implies_hpar(2))
proposition is_weak_composition:
shows "weak_composition V H"
proof -
have 1: "(\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu> \<noteq> null)
= (\<lambda>\<mu>\<nu>. arr (fst \<mu>\<nu>) \<and> arr (snd \<mu>\<nu>) \<and> src (fst \<mu>\<nu>) = trg (snd \<mu>\<nu>))"
using hseq_char' by auto
interpret VoV: subcategory VxV.comp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu> \<noteq> null\<close>
using 1 VV.subcategory_axioms by simp
interpret H: "functor" VoV.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
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 \<open>
It can be shown that \<open>arr ((\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>)) \<Longrightarrow> (\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>) = (\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)\<close>.
However, we do not have \<open>arr ((\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)) \<Longrightarrow> (\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>) = (\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)\<close>,
because it does not follow from \<open>arr ((\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>))\<close> that \<open>dom \<nu> = cod \<mu>\<close>
and \<open>dom \<tau> = cod \<sigma>\<close>, only that \<open>dom \<nu> \<star> dom \<tau> = cod \<mu> \<star> cod \<sigma>\<close>.
So we don't get interchange unconditionally.
\<close>
lemma interchange:
assumes "seq \<nu> \<mu>" and "seq \<tau> \<sigma>"
shows "(\<nu> \<cdot> \<mu>) \<star> (\<tau> \<cdot> \<sigma>) = (\<nu> \<star> \<tau>) \<cdot> (\<mu> \<star> \<sigma>)"
using assms interchange by simp
lemma whisker_right:
assumes "ide f" and "seq \<nu> \<mu>"
shows "(\<nu> \<cdot> \<mu>) \<star> f = (\<nu> \<star> f) \<cdot> (\<mu> \<star> f)"
using assms whisker_right by simp
lemma whisker_left:
assumes "ide f" and "seq \<nu> \<mu>"
shows "f \<star> (\<nu> \<cdot> \<mu>) = (f \<star> \<nu>) \<cdot> (f \<star> \<mu>)"
using assms whisker_left by simp
lemma inverse_arrows_hcomp:
assumes "iso \<mu>" and "iso \<nu>" and "src \<nu> = trg \<mu>"
shows "inverse_arrows (\<nu> \<star> \<mu>) (inv \<nu> \<star> inv \<mu>)"
proof -
show "inverse_arrows (\<nu> \<star> \<mu>) (inv \<nu> \<star> inv \<mu>)"
proof
show "ide ((inv \<nu> \<star> inv \<mu>) \<cdot> (\<nu> \<star> \<mu>))"
proof -
have "(inv \<nu> \<star> inv \<mu>) \<cdot> (\<nu> \<star> \<mu>) = dom \<nu> \<star> dom \<mu>"
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 ((\<nu> \<star> \<mu>) \<cdot> (inv \<nu> \<star> inv \<mu>))"
proof -
have "(\<nu> \<star> \<mu>) \<cdot> (inv \<nu> \<star> inv \<mu>) = cod \<nu> \<star> cod \<mu>"
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 \<mu>" and "iso \<nu>" and "src \<nu> = trg \<mu>"
shows "iso (\<nu> \<star> \<mu>)"
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 \<mu> f g" and "iso_in_hom \<nu> h k" and "src \<nu> = trg \<mu>"
shows "iso_in_hom (\<nu> \<star> \<mu>) (h \<star> f) (k \<star> 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 \<cong> g"
shows "ide f" and "ide g"
using assms isomorphic_def by auto
lemma hcomp_ide_isomorphic:
assumes "ide f" and "g \<cong> h" and "src f = trg g"
shows "f \<star> g \<cong> f \<star> h"
proof -
obtain \<mu> where \<mu>: "iso \<mu> \<and> \<guillemotleft>\<mu> : g \<Rightarrow> h\<guillemotright>"
using assms isomorphic_def by auto
have "iso (f \<star> \<mu>) \<and> \<guillemotleft>f \<star> \<mu> : f \<star> g \<Rightarrow> f \<star> h\<guillemotright>"
using assms \<mu> iso_hcomp by auto
thus ?thesis
using isomorphic_def by auto
qed
lemma hcomp_isomorphic_ide:
assumes "f \<cong> g" and "ide h" and "src f = trg h"
shows "f \<star> h \<cong> g \<star> h"
proof -
obtain \<mu> where \<mu>: "iso \<mu> \<and> \<guillemotleft>\<mu> : f \<Rightarrow> g\<guillemotright>"
using assms isomorphic_def by auto
have "iso (\<mu> \<star> h) \<and> \<guillemotleft>\<mu> \<star> h : f \<star> h \<Rightarrow> g \<star> h\<guillemotright>"
using assms \<mu> iso_hcomp by auto
thus ?thesis
using isomorphic_def by auto
qed
lemma isomorphic_implies_hpar:
assumes "f \<cong> 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 \<nu>" and "iso \<mu>" and "src \<nu> = trg \<mu>"
shows "inv (\<nu> \<star> \<mu>) = inv \<nu> \<star> inv \<mu>"
using assms inverse_arrow_unique [of "\<nu> \<star> \<mu>"] inv_is_inverse inverse_arrows_hcomp
by auto
text \<open>
The following define the two ways of using horizontal composition to compose three arrows.
\<close>
definition HoHV
where "HoHV \<mu> \<equiv> if VVV.arr \<mu> then (fst \<mu> \<star> fst (snd \<mu>)) \<star> snd (snd \<mu>) else null"
definition HoVH
where "HoVH \<mu> \<equiv> if VVV.arr \<mu> then fst \<mu> \<star> fst (snd \<mu>) \<star> snd (snd \<mu>) 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
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C 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 \<cdot> HoHV f"
proof -
have "VxVxV.comp g f =
(fst g \<cdot> fst f, fst (snd g) \<cdot> fst (snd f), snd (snd g) \<cdot> snd (snd f))"
- using fg VVV.seq_char VVV.arr_char VV.arr_char VxVxV.comp_char VxV.comp_char
+ using fg VVV.seq_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VxVxV.comp_char VxV.comp_char
by (metis (no_types, lifting) VxV.seqE VxVxV.seqE)
hence "HoHV (VVV.comp g f) =
(fst g \<cdot> fst f \<star> fst (snd g) \<cdot> fst (snd f)) \<star> snd (snd g) \<cdot> snd (snd f)"
using HoHV_def VVV.comp_simp fg by auto
also have "... = ((fst g \<star> fst (snd g)) \<star> snd (snd g)) \<cdot>
((fst f \<star> fst (snd f)) \<star> snd (snd f))"
- using fg VVV.seq_char VVV.arr_char VV.arr_char interchange
+ using fg VVV.seq_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C interchange
by (metis (no_types, lifting) VxV.seqE VxVxV.seqE hseqI' src_vcomp trg_vcomp)
also have "... = HoHV g \<cdot> 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
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C 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 \<cdot> HoVH f"
proof -
have "VxVxV.comp g f =
(fst g \<cdot> fst f, fst (snd g) \<cdot> fst (snd f), snd (snd g) \<cdot> snd (snd f))"
- using fg VVV.seq_char VVV.arr_char VV.arr_char VxVxV.comp_char VxV.comp_char
+ using fg VVV.seq_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VxVxV.comp_char VxV.comp_char
by (metis (no_types, lifting) VxV.seqE VxVxV.seqE)
hence "HoVH (VVV.comp g f) =
fst g \<cdot> fst f \<star> fst (snd g) \<cdot> fst (snd f) \<star> snd (snd g) \<cdot> snd (snd f)"
using HoVH_def VVV.comp_simp fg by auto
also have "... = (fst g \<star> fst (snd g) \<star> snd (snd g)) \<cdot>
(fst f \<star> fst (snd f) \<star> snd (snd f))"
- using fg VVV.seq_char VVV.arr_char VV.arr_char interchange
+ using fg VVV.seq_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C interchange
by (metis (no_types, lifting) VxV.seqE VxVxV.seqE hseqI' src_vcomp trg_vcomp)
also have "... = HoVH g \<cdot> HoVH f"
- using fg VVV.seq_char VVV.arr_char HoVH_def VVV.comp_char VV.arr_char
+ using fg VVV.seq_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C HoVH_def VVV.comp_char VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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 \<open>
The following define horizontal composition of an arrow on the left by its target
and on the right by its source.
\<close>
abbreviation L
where "L \<equiv> \<lambda>\<mu>. if arr \<mu> then trg \<mu> \<star> \<mu> else null"
abbreviation R
where "R \<equiv> \<lambda>\<mu>. if arr \<mu> then \<mu> \<star> src \<mu> 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,2959 +1,2959 @@
(* Title: Pseudofunctor
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Pseudofunctors"
theory Pseudofunctor
imports MonoidalCategory.MonoidalFunctor Bicategory Subbicategory InternalEquivalence Coherence
begin
text \<open>
The traditional definition of a pseudofunctor \<open>F : C \<rightarrow> D\<close> between bicategories \<open>C\<close> and \<open>D\<close>
is in terms of two maps: an ``object map'' \<open>F\<^sub>o\<close> that takes objects of \<open>C\<close> to objects of \<open>D\<close>
and an ``arrow map'' \<open>F\<^sub>a\<close> that assigns to each pair of objects \<open>a\<close> and \<open>b\<close> of \<open>C\<close>
a functor \<open>F\<^sub>a a b\<close> from the hom-category \<open>hom\<^sub>C a b\<close> to the hom-category \<open>hom\<^sub>D (F\<^sub>o a) (F\<^sub>o b)\<close>.
In addition, there is assigned to each object \<open>a\<close> of \<open>C\<close> an invertible 2-cell
\<open>\<guillemotleft>\<Psi> a : F\<^sub>o a \<Rightarrow>\<^sub>D (F\<^sub>a a a) a\<guillemotright>\<close>, and to each pair \<open>(f, g)\<close> of composable 1-cells of C there
is assigned an invertible 2-cell \<open>\<guillemotleft>\<Phi> (f, g) : F g \<star> F f \<Rightarrow> F (g \<star> f)\<guillemotright>\<close>, 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 \<open>F : C \<rightarrow> D\<close> is a functor from
\<open>C\<close> to \<open>D\<close>, when these are regarded as ordinary categories with respect to vertical
composition. Next, \<open>F\<close> 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 \<open>(f, g)\<close> of \<open>C\<close>, a ``compositor''
\<open>\<guillemotleft>\<Phi> (f, g) : F g \<star> F f \<Rightarrow> F (g \<star> f)\<guillemotright>\<close> in \<open>D\<close>, such that the \<open>\<Phi> (f, g)\<close> 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 \<open>\<a>\<^sub>C[f, g, h]\<close> in \<open>C\<close>, its image \<open>F \<a>\<^sub>C[f, g, h]\<close>,
the associator \<open>\<a>\<^sub>D[F f, F g, F h]\<close> in \<open>D\<close> and compositors involving \<open>f\<close>, \<open>g\<close>, and \<open>h\<close>.
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 \<open>F \<i>\<^sub>C[a]\<close> and \<open>\<i>\<^sub>D[F a]\<close> are isomorphic in \<open>D\<close> for each object \<open>a\<close> of \<open>C\<close>,
for there is then a canonical way to obtain, for each \<open>a\<close>, an isomorphism
\<open>\<guillemotleft>\<Psi> a : src (F a) \<rightarrow> F a\<guillemotright>\<close> that satisfies the usual coherence conditions relating the
unitors and the associators. Note that the map \<open>a \<mapsto> src (F a)\<close> amounts to the traditional
``object map'' \<open>F\<^sub>o\<close>, so that this becomes a derived notion, rather than a primitive one.
\<close>
subsection "Weak Arrows of Homs"
text \<open>
We begin with a locale that defines a functor between ``horizontal homs'' that preserves
source and target up to isomorphism.
\<close>
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 "\<cdot>\<^sub>C" 55)
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd" +
assumes weakly_preserves_src: "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> D.isomorphic (F (src\<^sub>C \<mu>)) (src\<^sub>D (F \<mu>))"
and weakly_preserves_trg: "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> D.isomorphic (F (trg\<^sub>C \<mu>)) (trg\<^sub>D (F \<mu>))"
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 \<mu> \<nu> \<equiv> C.arr \<mu> \<and> C.arr \<nu> \<and> src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
abbreviation (input) hseq\<^sub>D
where "hseq\<^sub>D \<mu> \<nu> \<equiv> D.arr \<mu> \<and> D.arr \<nu> \<and> src\<^sub>D \<mu> = trg\<^sub>D \<nu>"
lemma preserves_hseq:
assumes "hseq\<^sub>C \<mu> \<nu>"
shows "hseq\<^sub>D (F \<mu>) (F \<nu>)"
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 \<open>
Though \<open>F\<close> 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 \<open>F\<close>.
\<close>
definition map\<^sub>0
where "map\<^sub>0 a \<equiv> 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 \<mu>"
shows "src\<^sub>D (F \<mu>) = map\<^sub>0 (src\<^sub>C \<mu>)"
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 \<mu>"
shows "trg\<^sub>D (F \<mu>) = map\<^sub>0 (trg\<^sub>C \<mu>)"
using assms map\<^sub>0_def preserves_hseq C.src_trg C.trg.preserves_arr by presburger
lemma preserves_hhom [intro]:
assumes "C.arr \<mu>"
shows "D.in_hhom (F \<mu>) (map\<^sub>0 (src\<^sub>C \<mu>)) (map\<^sub>0 (trg\<^sub>C \<mu>))"
using assms by simp
text \<open>
We define here the lifting of \<open>F\<close> to a functor \<open>FF: CC \<rightarrow> DD\<close>.
We need this to define the domains and codomains of the compositors.
\<close>
definition FF
where "FF \<equiv> \<lambda>\<mu>\<nu>. if C.VV.arr \<mu>\<nu> then (F (fst \<mu>\<nu>), F (snd \<mu>\<nu>)) else D.VV.null"
sublocale FF: "functor" C.VV.comp D.VV.comp FF
proof -
have 1: "\<And>\<mu>\<nu>. C.VV.arr \<mu>\<nu> \<Longrightarrow> D.VV.arr (FF \<mu>\<nu>)"
- unfolding FF_def using C.VV.arr_char D.VV.arr_char preserves_hseq by simp
+ unfolding FF_def using C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C preserves_hseq by simp
show "functor C.VV.comp D.VV.comp FF"
proof
fix \<mu>\<nu>
show "\<not> C.VV.arr \<mu>\<nu> \<Longrightarrow> FF \<mu>\<nu> = D.VV.null"
using FF_def by simp
show "C.VV.arr \<mu>\<nu> \<Longrightarrow> D.VV.arr (FF \<mu>\<nu>)"
using 1 by simp
assume \<mu>\<nu>: "C.VV.arr \<mu>\<nu>"
show "D.VV.dom (FF \<mu>\<nu>) = FF (C.VV.dom \<mu>\<nu>)"
- using \<mu>\<nu> 1 FF_def C.VV.arr_char D.VV.arr_char C.VV.dom_simp D.VV.dom_simp
+ using \<mu>\<nu> 1 FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp D.VV.dom_simp
by simp
show "D.VV.cod (FF \<mu>\<nu>) = FF (C.VV.cod \<mu>\<nu>)"
- using \<mu>\<nu> 1 FF_def C.VV.arr_char D.VV.arr_char C.VV.cod_simp D.VV.cod_simp
+ using \<mu>\<nu> 1 FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_simp D.VV.cod_simp
by simp
next
fix \<mu>\<nu> \<tau>\<pi>
assume 2: "C.VV.seq \<mu>\<nu> \<tau>\<pi>"
show "FF (C.VV.comp \<mu>\<nu> \<tau>\<pi>) = D.VV.comp (FF \<mu>\<nu>) (FF \<tau>\<pi>)"
proof -
have "FF (C.VV.comp \<mu>\<nu> \<tau>\<pi>) = (F (fst \<mu>\<nu>) \<cdot>\<^sub>D F (fst \<tau>\<pi>), F (snd \<mu>\<nu>) \<cdot>\<^sub>D F (snd \<tau>\<pi>))"
- 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
+ using 1 2 FF_def C.VV.comp_char C.VxV.comp_char C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis (no_types, lifting) C.VV.seq_char\<^sub>S\<^sub>b\<^sub>C C.VxV.seqE fst_conv
as_nat_trans.preserves_comp_2 snd_conv)
also have "... = D.VV.comp (FF \<mu>\<nu>) (FF \<tau>\<pi>)"
- 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
+ using 1 2 FF_def D.VV.comp_char D.VxV.comp_char C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.seq_char\<^sub>S\<^sub>b\<^sub>C 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 \<open>
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 \<open>https://ncatlab.org/nlab/show/pseudofunctor\<close>}),
but that is not workable here because we need a name that does not clash with that
used for an ordinary functor between categories.
\<close>
locale pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^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 \<open>\<lambda>\<mu>\<nu>. H\<^sub>C (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> F +
H\<^sub>DoFF: composite_functor C.VV.comp D.VV.comp V\<^sub>D
FF \<open>\<lambda>\<mu>\<nu>. H\<^sub>D (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> +
\<Phi>: natural_isomorphism C.VV.comp V\<^sub>D H\<^sub>DoFF.map FoH\<^sub>C.map \<Phi>
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi> :: "'c * 'c \<Rightarrow> 'd" +
assumes assoc_coherence:
"\<lbrakk> 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 \<rbrakk> \<Longrightarrow>
F \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
begin
no_notation C.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
no_notation D.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>D _\<guillemotright>")
notation C.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
notation C.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>C _\<guillemotright>")
notation D.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>D _\<guillemotright>")
notation D.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>D _\<guillemotright>")
notation C.lunit ("\<l>\<^sub>C[_]")
notation C.runit ("\<r>\<^sub>C[_]")
notation C.lunit' ("\<l>\<^sub>C\<^sup>-\<^sup>1[_]")
notation C.runit' ("\<r>\<^sub>C\<^sup>-\<^sup>1[_]")
notation C.\<a>' ("\<a>\<^sub>C\<^sup>-\<^sup>1[_, _, _]")
notation D.lunit ("\<l>\<^sub>D[_]")
notation D.runit ("\<r>\<^sub>D[_]")
notation D.lunit' ("\<l>\<^sub>D\<^sup>-\<^sup>1[_]")
notation D.runit' ("\<r>\<^sub>D\<^sup>-\<^sup>1[_]")
notation D.\<a>' ("\<a>\<^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 "\<guillemotleft>\<Phi> (a, b) : map\<^sub>0 (src\<^sub>C b) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C a)\<guillemotright>"
and "\<guillemotleft>\<Phi> (a, b) : F a \<star>\<^sub>D F b \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C b)\<guillemotright>"
proof -
show "\<guillemotleft>\<Phi> (a, b) : F a \<star>\<^sub>D F b \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C b)\<guillemotright>"
- using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by auto
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def by auto
thus "\<guillemotleft>\<Phi> (a, b) : map\<^sub>0 (src\<^sub>C b) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C a)\<guillemotright>"
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 (\<Phi> (f, g))"
and "src\<^sub>D (\<Phi> (f, g)) = src\<^sub>D (F g)" and "trg\<^sub>D (\<Phi> (f, g)) = trg\<^sub>D (F f)"
and "D.dom (\<Phi> (f, g)) = F f \<star>\<^sub>D F g" and "D.cod (\<Phi> (f, g)) = F (f \<star>\<^sub>C g)"
using assms cmp_in_hom by simp_all blast+
lemma cmp_in_hom':
assumes "C.arr \<mu>" and "C.arr \<nu>" and "src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
shows "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : map\<^sub>0 (src\<^sub>C \<nu>) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>"
and "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>) \<Rightarrow>\<^sub>D F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)\<guillemotright>"
proof -
show "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>) \<Rightarrow>\<^sub>D F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)\<guillemotright>"
- using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by auto
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def by auto
thus "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : map\<^sub>0 (src\<^sub>C \<nu>) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>"
using assms D.vconn_implies_hpar by auto
qed
lemma cmp_simps':
assumes "C.arr \<mu>" and "C.arr \<nu>" and "src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
shows "D.arr (\<Phi> (\<mu>, \<nu>))"
and "src\<^sub>D (\<Phi> (\<mu>, \<nu>)) = map\<^sub>0 (src\<^sub>C \<nu>)" and "trg\<^sub>D (\<Phi> (\<mu>, \<nu>)) = map\<^sub>0 (trg\<^sub>C \<mu>)"
and "D.dom (\<Phi> (\<mu>, \<nu>)) = F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>)"
and "D.cod (\<Phi> (\<mu>, \<nu>)) = F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)"
using assms cmp_in_hom' by 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 (\<Phi> (f, g))"
- using assms C.VV.ide_char C.VV.arr_char by simp
+ using assms C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<star>\<^sub>D F g) (F (f \<star>\<^sub>C g))"
using assms D.isomorphic_def by auto
end
context pseudofunctor
begin
text \<open>
The following defines the image of the unit isomorphism \<open>\<i>\<^sub>C[a]\<close> under \<open>F\<close>.
We will use \<open>(F a, \<i>[a])\<close> as an ``alternate unit'', to substitute for
\<open>(src\<^sub>D (F a), \<i>\<^sub>D[src\<^sub>D (F a)])\<close>.
\<close>
abbreviation (input) \<i> ("\<i>[_]")
where "\<i>[a] \<equiv> F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a)"
lemma \<i>_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a) : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<i>[a] : F a \<star>\<^sub>D F a \<Rightarrow>\<^sub>D F a\<guillemotright>"
proof (unfold map\<^sub>0_def)
show "\<guillemotleft>F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a) : F a \<star>\<^sub>D F a \<Rightarrow>\<^sub>D F a\<guillemotright>"
using assms preserves_hom cmp_in_hom
by (intro D.comp_in_homI, auto)
show "\<guillemotleft>F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a) : src\<^sub>D (F a) \<rightarrow>\<^sub>D src\<^sub>D (F a)\<guillemotright>"
- using assms C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
by (intro D.vcomp_in_hhom D.seqI, auto)
qed
lemma \<i>_simps [simp]:
assumes "C.obj a"
shows "D.arr (\<i> a)"
and "src\<^sub>D \<i>[a] = map\<^sub>0 a" and "trg\<^sub>D \<i>[a] = map\<^sub>0 a"
and "D.dom \<i>[a] = F a \<star>\<^sub>D F a" and "D.cod \<i>[a] = F a"
using assms \<i>_in_hom by auto
lemma iso_\<i>:
assumes "C.obj a"
shows "D.iso \<i>[a]"
using assms C.iso_unit C.obj_self_composable(1) C.seq_if_composable
by (meson C.objE D.isos_compose \<i>_simps(1) cmp_components_are_iso preserves_iso)
text \<open>
If \<open>a\<close> is an object of \<open>C\<close> and we have an isomorphism \<open>\<guillemotleft>\<Phi> (a, a) : F a \<star>\<^sub>D F a \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C a)\<guillemotright>\<close>,
then there is a canonical way to define a compatible isomorphism \<open>\<guillemotleft>\<Psi> a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>\<close>.
Specifically, we take \<open>\<Psi> a\<close> to be the unique isomorphism \<open>\<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>\<close> such that
\<open>\<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)\<close>.
\<close>
definition unit
where "unit a \<equiv> THE \<psi>. \<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<psi> \<and>
\<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)"
lemma unit_char:
assumes "C.obj a"
shows "\<guillemotleft>unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>" and "D.iso (unit a)"
and "unit a \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (unit a \<star>\<^sub>D unit a)"
and "\<exists>!\<psi>. \<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<psi> \<and> \<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)"
proof -
let ?P = "\<lambda>\<psi>. \<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<psi> \<and> \<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)"
show "\<exists>!\<psi>. ?P \<psi>"
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 \<Phi>.preserves_hom \<i>_in_hom iso_\<i> by simp
qed
hence 1: "?P (unit a)"
using assms unit_def the1I2 [of ?P ?P] by simp
show "\<guillemotleft>unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>" using 1 by simp
show "D.iso (unit a)" using 1 by simp
show "unit a \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (unit a \<star>\<^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"
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 "\<guillemotleft>unit a : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>"
using assms by auto
lemma unit_eqI:
assumes "C.obj a" and "\<guillemotleft>\<mu>: map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>" and "D.iso \<mu>"
and "\<mu> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i> a \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D \<mu>)"
shows "\<mu> = unit a"
using assms unit_def unit_char
the1_equality [of "\<lambda>\<mu>. \<guillemotleft>\<mu> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<mu> \<and>
\<mu> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D \<mu>)" \<mu>]
by simp
text \<open>
The following defines the unique isomorphism satisfying the characteristic conditions
for the left unitor \<open>\<l>\<^sub>D[trg\<^sub>D (F f)]\<close>, but using the ``alternate unit'' \<open>\<i>[trg\<^sub>C f]\<close>
instead of \<open>\<i>\<^sub>D[trg\<^sub>D (F f)]\<close>, which is used to define \<open>\<l>\<^sub>D[trg\<^sub>D (F f)]\<close>.
\<close>
definition lF
where "lF f \<equiv> THE \<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<and>
F (trg\<^sub>C f) \<star>\<^sub>D \<mu> =(\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^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 "\<guillemotleft>lF f : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright>"
and "F (trg\<^sub>C f) \<star>\<^sub>D lF f = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<and>
F (trg\<^sub>C f) \<star>\<^sub>D \<mu> = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
proof -
let ?P = "\<lambda>\<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<and>
F (trg\<^sub>C f) \<star>\<^sub>D \<mu> = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
show "\<exists>!\<mu>. ?P \<mu>"
proof -
interpret Df: prebicategory \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D
using D.is_prebicategory by simp
interpret S: subcategory \<open>(\<cdot>\<^sub>D)\<close> \<open>Df.left (F (trg\<^sub>C f))\<close>
using assms Df.left_hom_is_subcategory by simp
interpret Df: left_hom \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<open>F (trg\<^sub>C f)\<close>
using assms D.weak_unit_char
by unfold_locales simp
interpret Df: left_hom_with_unit \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D \<open>\<i>[trg\<^sub>C f]\<close> \<open>F (trg\<^sub>C f)\<close>
using assms \<i>_in_hom iso_\<i> D.weak_unit_char(1) assms weakly_preserves_trg
by unfold_locales auto
have "\<exists>!\<mu>. \<guillemotleft>\<mu> : Df.L (F f) \<Rightarrow>\<^sub>S F f\<guillemotright> \<and>
Df.L \<mu> = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>S \<a>\<^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
+ using assms Df.lunit_char(3) S.ide_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
moreover have "Df.L (F f) = F (trg\<^sub>C f) \<star>\<^sub>D F f"
using assms by (simp add: Df.H\<^sub>L_def)
moreover have "\<And>\<mu>. Df.L \<mu> = F (trg\<^sub>C f) \<star>\<^sub>D \<mu>"
using Df.H\<^sub>L_def by simp
moreover have "(\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>S \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f] =
(\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
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.\<iota>_in_hom Df.arr_\<omega> S.comp_char S.in_hom_char
+ D.vconn_implies_hpar(4) Df.\<iota>_in_hom Df.arr_\<omega> S.comp_char S.in_hom_char\<^sub>S\<^sub>b\<^sub>C
calculation(1,3))
moreover have "\<And>\<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<longleftrightarrow>
\<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>S F f\<guillemotright>"
- using assms S.in_hom_char S.arr_char
+ using assms S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C
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 "\<guillemotleft>lF f : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright>"
using 1 by simp
show "F (trg\<^sub>C f) \<star>\<^sub>D lF f = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^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) \<star>\<^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 \<open>
\sloppypar
The next two lemmas generalize the eponymous results from
@{theory MonoidalCategory.MonoidalFunctor}. See the proofs of those results for diagrams.
\<close>
lemma lunit_coherence1:
assumes "C.ide f"
shows "\<l>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (unit (trg\<^sub>C f) \<star>\<^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 \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f) = \<l>\<^sub>D[F f]"
proof -
have "D.par (lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f)) \<l>\<^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 \<star>\<^sub>D (lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f)) = map\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f]"
proof -
have "map\<^sub>0 ?b \<star>\<^sub>D (lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f)) =
(map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (map\<^sub>0 ?b \<star>\<^sub>D unit ?b \<star>\<^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 \<star>\<^sub>D F f"]
by auto
also have "... = (map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D
(D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
proof -
have "(D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f) =
D.inv (unit ?b) \<cdot>\<^sub>D unit ?b \<star>\<^sub>D F ?b \<cdot>\<^sub>D unit ?b \<star>\<^sub>D F f \<cdot>\<^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 \<star>\<^sub>D F f" "unit ?b \<star>\<^sub>D F f"]
by simp
also have "... = map\<^sub>0 ?b \<star>\<^sub>D unit ?b \<star>\<^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) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
proof -
have "(map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) =
(D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
proof -
have "(map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) =
D.inv (unit ?b) \<star>\<^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 \<star>\<^sub>D F f"]
by simp
also have "... = D.inv (unit ?b) \<cdot>\<^sub>D F ?b \<star>\<^sub>D F f \<cdot>\<^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) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F ?b \<star>\<^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) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f] \<cdot>\<^sub>D
(unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
using assms unit_char(1-2) lF_char(2) D.comp_assoc by auto
also have "... = ((D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D
((unit ?b \<star>\<^sub>D unit ?b) \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<a>\<^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: \<open>trg\<^sub>D (F f) = map\<^sub>0 (trg\<^sub>C f)\<close>)
also have "... = (\<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 ?b, map\<^sub>0 ?b, F f]"
proof -
have "((D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D ((unit ?b \<star>\<^sub>D unit ?b) \<star>\<^sub>D F f)) =
\<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f"
proof -
have "((D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D ((unit ?b \<star>\<^sub>D unit ?b) \<star>\<^sub>D F f)) =
D.inv (unit ?b) \<cdot>\<^sub>D unit ?b \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^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 "... = \<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f"
proof -
have "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D unit (trg\<^sub>C f)) \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 ?b] = \<i>\<^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 \<star>\<^sub>D \<l>\<^sub>D[F f]"
using assms D.lunit_char [of "F f"] \<open>trg\<^sub>D (F f) = map\<^sub>0 ?b\<close> 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 \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f)"
proof -
let ?b = "trg\<^sub>C f"
have "D.par (F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (?b, f)) (lF f)"
using assms cmp_simps'(1) cmp_simps(4-5) by force
moreover have "F ?b \<star>\<^sub>D F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (?b, f) = F ?b \<star>\<^sub>D lF f"
proof -
have "F ?b \<star>\<^sub>D F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (?b, f) = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D \<Phi> (?b, f))"
using assms cmp_in_hom D.whisker_left [of "F ?b" "F \<l>\<^sub>C[f]" "\<Phi> (?b, f)"]
by (simp add: calculation)
also have "... = F ?b \<star>\<^sub>D lF f"
proof -
have "(F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D \<Phi> (?b, f))
= (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D
\<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
proof -
have 1: "D.seq (F \<a>\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f])
(\<Phi> (trg\<^sub>C f \<star>\<^sub>C trg\<^sub>C f, f) \<cdot>\<^sub>D (\<Phi> (trg\<^sub>C f, trg\<^sub>C f) \<star>\<^sub>D F f))"
using assms by fastforce
hence 2: "D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f)
= (F ?b \<star>\<^sub>D \<Phi> (?b, f)) \<cdot>\<^sub>D \<a>\<^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 \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f)"
"\<Phi> (?b, ?b \<star>\<^sub>C f)"
"(F ?b \<star>\<^sub>D \<Phi> (?b, f)) \<cdot>\<^sub>D \<a>\<^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 \<star>\<^sub>D \<Phi> (?b, f)
= (D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
proof -
have "D.seq (D.inv (\<Phi> (trg\<^sub>C f, trg\<^sub>C f \<star>\<^sub>C f)))
(F \<a>\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f \<star>\<^sub>C trg\<^sub>C f, f) \<cdot>\<^sub>D
(\<Phi> (trg\<^sub>C f, trg\<^sub>C f) \<star>\<^sub>D F f))"
using assms 1 D.hseq_char by auto
moreover have "(F (trg\<^sub>C f) \<star>\<^sub>D \<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F (trg\<^sub>C f), F (trg\<^sub>C f), F f] =
D.inv (\<Phi> (trg\<^sub>C f, trg\<^sub>C f \<star>\<^sub>C f)) \<cdot>\<^sub>D
F \<a>\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f \<star>\<^sub>C trg\<^sub>C f, f) \<cdot>\<^sub>D
(\<Phi> (trg\<^sub>C f, trg\<^sub>C f) \<star>\<^sub>D F f)"
using assms 2 by simp
ultimately show ?thesis
using assms
D.invert_side_of_triangle(2)
[of "D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f)"
"F ?b \<star>\<^sub>D \<Phi> (?b, f)" "\<a>\<^sub>D[F ?b, F ?b, F f]"]
by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
(D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D F (\<i>\<^sub>C[?b] \<star>\<^sub>C f)) \<cdot>\<^sub>D
\<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
proof -
have 1: "F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) = F (\<i>\<^sub>C[?b] \<star>\<^sub>C f) \<cdot>\<^sub>D D.inv (F \<a>\<^sub>C[?b, ?b, f])"
using assms C.lunit_char(1-2) C.unit_in_hom preserves_inv by auto
have "F \<a>\<^sub>C[?b, ?b, f] = D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D F (\<i>\<^sub>C[?b] \<star>\<^sub>C f)"
proof -
have "F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D D.inv (F (\<i>\<^sub>C[?b] \<star>\<^sub>C f)) =
D.inv (F (\<i>\<^sub>C[?b] \<star>\<^sub>C f) \<cdot>\<^sub>D D.inv (F \<a>\<^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 \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D (F (\<i>\<^sub>C[?b] \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f)) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
using D.comp_assoc by auto
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D (\<Phi> (?b, f) \<cdot>\<^sub>D (F \<i>\<^sub>C[?b] \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
- using assms \<Phi>.naturality [of "(\<i>\<^sub>C[?b], f)"] FF_def C.VV.arr_char C.VV.cod_char
- C.VV.dom_char
+ using assms \<Phi>.naturality [of "(\<i>\<^sub>C[?b], f)"] FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f) \<cdot>\<^sub>D
((F \<i>\<^sub>C[?b] \<star>\<^sub>D F f)) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
using D.comp_assoc by auto
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^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 \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D
(D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f)) \<cdot>\<^sub>D
(F ?b \<star>\<^sub>D lF f)"
using D.comp_assoc assms lF_char(2) by presburger
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
proof -
have "D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) =
D.inv (((F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f))) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^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 (\<Phi> (?b, f)) \<cdot>\<^sub>D F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
proof -
have 1: "\<Phi> (?b, f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) = F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f)"
- using assms \<Phi>.naturality [of "(?b, \<l>\<^sub>C[f])"] FF_def C.VV.arr_char
- C.VV.cod_char D.VV.null_char C.VV.dom_simp
+ using assms \<Phi>.naturality [of "(?b, \<l>\<^sub>C[f])"] FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.VV.null_char C.VV.dom_simp
by simp
have "(F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) =
D.inv (\<Phi> (?b, f)) \<cdot>\<^sub>D F (?b \<star>\<^sub>C \<l>\<^sub>C[f])"
proof -
have "D.seq (\<Phi> (?b, f)) (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f])"
using assms cmp_in_hom(2) [of ?b f] by auto
moreover have "D.iso (\<Phi> (?b, f)) \<and> D.iso (\<Phi> (?b, ?b \<star>\<^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 \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D \<Phi> (?b, f)"
proof -
have "D.iso (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
- using assms D.isos_compose C.VV.arr_char C.iso_lunit C.VV.dom_simp
+ using assms D.isos_compose C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.iso_lunit C.VV.dom_simp
C.VV.cod_simp
by simp
moreover have "D.iso (D.inv (\<Phi> (?b, f)))"
using assms by simp
moreover have "D.seq (D.inv (\<Phi> (?b, f))) (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
- using assms C.VV.arr_char C.VV.dom_simp C.VV.cod_simp by simp
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f)"
using D.comp_assoc D.inv_comp assms cmp_simps'(1) cmp_simps(5) by force
finally have "D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f])
= D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f)"
by blast
thus ?thesis by argo
qed
also have "... = ((F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f])) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
using D.comp_assoc by simp
also have "... = F ?b \<star>\<^sub>D lF f"
using assms D.comp_arr_inv' [of "F ?b \<star>\<^sub>D F \<l>\<^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 "\<l>\<^sub>D[F f] = F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
proof -
have "\<l>\<^sub>D[F f] = (F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
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 D.comp_assoc by simp
qed
text \<open>
We postpone proving the dual version of this result until after we have developed
the notion of the ``op bicategory'' in the next section.
\<close>
end
subsection "Pseudofunctors and Opposite Bicategories"
text \<open>
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.
\<close>
locale op_bicategory =
B: bicategory V H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
for V :: "'a comp" (infixr "\<cdot>" 55)
and H\<^sub>B :: "'a comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'a \<Rightarrow> 'a" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
begin
abbreviation H (infixr "\<star>" 53)
where "H f g \<equiv> H\<^sub>B g f"
abbreviation \<i> ("\<i>[_]")
where "\<i> \<equiv> \<i>\<^sub>B"
abbreviation src
where "src \<equiv> trg\<^sub>B"
abbreviation trg
where "trg \<equiv> src\<^sub>B"
interpretation horizontal_homs V src trg
by (unfold_locales, auto)
interpretation H: "functor" VV.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
- using VV.arr_char VV.dom_simp VV.cod_simp
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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
+ using VV.comp_char VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu>\<nu>\<tau> \<equiv> if B.VVV.arr \<mu>\<nu>\<tau> then
(snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)
else VVV.null"
abbreviation DN
where "DN \<mu>\<nu>\<tau> \<equiv> if VVV.arr \<mu>\<nu>\<tau> then
(snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)
else B.VVV.null"
lemma VVV_arr_char:
shows "VVV.arr \<mu>\<nu>\<tau> \<longleftrightarrow> B.VVV.arr (DN \<mu>\<nu>\<tau>)"
- using VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char B.VVV.not_arr_null
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.not_arr_null
by auto
lemma VVV_ide_char:
shows "VVV.ide \<mu>\<nu>\<tau> \<longleftrightarrow> B.VVV.ide (DN \<mu>\<nu>\<tau>)"
proof -
have "VVV.ide \<mu>\<nu>\<tau> \<longleftrightarrow> VVV.arr \<mu>\<nu>\<tau> \<and> B.VxVxV.ide \<mu>\<nu>\<tau>"
- using VVV.ide_char by simp
+ using VVV.ide_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... \<longleftrightarrow> B.VVV.arr (DN \<mu>\<nu>\<tau>) \<and> B.VxVxV.ide (DN \<mu>\<nu>\<tau>)"
using VVV_arr_char B.VxVxV.ide_char by auto
also have "... \<longleftrightarrow> B.VVV.ide (DN \<mu>\<nu>\<tau>)"
- using B.VVV.ide_char [of "DN \<mu>\<nu>\<tau>"] by blast
+ using B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C [of "DN \<mu>\<nu>\<tau>"] by blast
finally show ?thesis by fast
qed
lemma VVV_dom_char:
shows "VVV.dom \<mu>\<nu>\<tau> = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
proof (cases "VVV.arr \<mu>\<nu>\<tau>")
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.dom \<mu>\<nu>\<tau> = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
using VVV.dom_def VVV.has_domain_iff_arr VVV_arr_char B.VVV.dom_null
by auto
show "VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.dom \<mu>\<nu>\<tau> = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
proof -
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have "VVV.dom \<mu>\<nu>\<tau> =
(B.dom (fst \<mu>\<nu>\<tau>), B.dom (fst (snd \<mu>\<nu>\<tau>)), B.dom (snd (snd \<mu>\<nu>\<tau>)))"
- using \<mu>\<nu>\<tau> VVV.dom_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
+ using \<mu>\<nu>\<tau> VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = UP (B.dom (snd (snd \<mu>\<nu>\<tau>)), B.dom (fst (snd \<mu>\<nu>\<tau>)), B.dom (fst \<mu>\<nu>\<tau>))"
- by (metis (no_types, lifting) B.VV.arrI B.VVV.arr_char B.arr_dom VV.arrE VVV.arrE
+ by (metis (no_types, lifting) B.VV.arrI\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.arr_dom VV.arrE VVV.arrE
\<mu>\<nu>\<tau> fst_conv snd_conv src_dom trg_dom)
also have "... = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
- using \<mu>\<nu>\<tau> B.VVV.dom_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char
+ using \<mu>\<nu>\<tau> B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally show ?thesis by blast
qed
qed
lemma VVV_cod_char:
shows "VVV.cod \<mu>\<nu>\<tau> = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
proof (cases "VVV.arr \<mu>\<nu>\<tau>")
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.cod \<mu>\<nu>\<tau> = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
using VVV.cod_def VVV.has_codomain_iff_arr VVV_arr_char B.VVV.cod_null
by auto
show "VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.cod \<mu>\<nu>\<tau> = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
proof -
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have "VVV.cod \<mu>\<nu>\<tau> = (B.cod (fst \<mu>\<nu>\<tau>), B.cod (fst (snd \<mu>\<nu>\<tau>)), B.cod (snd (snd \<mu>\<nu>\<tau>)))"
- using \<mu>\<nu>\<tau> VVV.cod_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
+ using \<mu>\<nu>\<tau> VVV.cod_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = UP (B.cod (snd (snd \<mu>\<nu>\<tau>)), B.cod (fst (snd \<mu>\<nu>\<tau>)), B.cod (fst \<mu>\<nu>\<tau>))"
- by (metis (no_types, lifting) B.VV.arrI B.VVV.arr_char B.arr_cod VV.arrE VVV.arrE
+ by (metis (no_types, lifting) B.VV.arrI\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.arr_cod VV.arrE VVV.arrE
\<mu>\<nu>\<tau> fst_conv snd_conv src_cod trg_cod)
also have "... = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
- using \<mu>\<nu>\<tau> B.VVV.cod_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char
+ using \<mu>\<nu>\<tau> B.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally show ?thesis by blast
qed
qed
lemma HoHV_char:
shows "HoHV \<mu>\<nu>\<tau> = B.HoVH (DN \<mu>\<nu>\<tau>)"
using HoHV_def B.HoVH_def VVV_arr_char by simp
lemma HoVH_char:
shows "HoVH \<mu>\<nu>\<tau> = B.HoHV (DN \<mu>\<nu>\<tau>)"
using HoVH_def B.HoHV_def VVV_arr_char by simp
definition \<a> ("\<a>[_, _, _]")
where "\<a>[\<mu>, \<nu>, \<tau>] \<equiv> B.\<alpha>' (DN (\<mu>, \<nu>, \<tau>))"
interpretation natural_isomorphism VVV.comp \<open>(\<cdot>)\<close> HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]\<close>
proof
fix \<mu>\<nu>\<tau>
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)] = B.null"
- using VVV.arr_char B.VVV.arr_char \<a>_def B.\<alpha>'.is_extensional by auto
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C \<a>_def B.\<alpha>'.is_extensional by auto
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
show "B.dom \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)] = HoHV (VVV.dom \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> \<a>_def HoHV_def B.\<alpha>'.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
+ using \<mu>\<nu>\<tau> \<a>_def HoHV_def B.\<alpha>'.preserves_dom VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.HoVH_def B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C
by auto
show "B.cod \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)] = HoVH (VVV.cod \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> \<a>_def HoVH_def B.\<alpha>'.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
+ using \<mu>\<nu>\<tau> \<a>_def HoVH_def B.\<alpha>'.preserves_cod VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.HoHV_def B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
by auto
show "HoVH \<mu>\<nu>\<tau> \<cdot>
\<a>[fst (VVV.dom \<mu>\<nu>\<tau>), fst (snd (VVV.dom \<mu>\<nu>\<tau>)), snd (snd (VVV.dom \<mu>\<nu>\<tau>))] =
\<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
proof -
have "HoVH \<mu>\<nu>\<tau> \<cdot>
\<a>[fst (VVV.dom \<mu>\<nu>\<tau>), fst (snd (VVV.dom \<mu>\<nu>\<tau>)), snd (snd (VVV.dom \<mu>\<nu>\<tau>))] =
HoVH \<mu>\<nu>\<tau> \<cdot> B.\<alpha>' (B.VVV.dom (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>))"
unfolding \<a>_def
- using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char VVV.dom_char B.VVV.arr_char B.VVV.dom_char
+ using \<mu>\<nu>\<tau> VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C
by auto
also have "... = B.\<alpha>' (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)"
using B.\<alpha>'.is_natural_1 VVV_arr_char \<mu>\<nu>\<tau> HoVH_char by presburger
also have "... = \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
using \<mu>\<nu>\<tau> \<a>_def by simp
finally show ?thesis by blast
qed
show "\<a>[fst (VVV.cod \<mu>\<nu>\<tau>), fst (snd (VVV.cod \<mu>\<nu>\<tau>)), snd (snd (VVV.cod \<mu>\<nu>\<tau>))] \<cdot>
HoHV \<mu>\<nu>\<tau> =
\<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))"
proof -
have "\<a>[fst (VVV.cod \<mu>\<nu>\<tau>), fst (snd (VVV.cod \<mu>\<nu>\<tau>)), snd (snd (VVV.cod \<mu>\<nu>\<tau>))] \<cdot>
HoHV \<mu>\<nu>\<tau> =
B.\<alpha>' (B.VVV.cod (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)) \<cdot> HoHV \<mu>\<nu>\<tau>"
unfolding \<a>_def
- using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char VVV.cod_char B.VVV.arr_char B.VVV.cod_char
+ using \<mu>\<nu>\<tau> VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.cod_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
by auto
also have "... = B.\<alpha>' (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)"
using B.\<alpha>'.is_natural_2 VVV_arr_char \<mu>\<nu>\<tau> HoHV_char by presburger
also have "... = \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
using \<mu>\<nu>\<tau> \<a>_def by simp
finally show ?thesis by blast
qed
next
fix \<mu>\<nu>\<tau>
assume \<mu>\<nu>\<tau>: "VVV.ide \<mu>\<nu>\<tau>"
show "B.iso \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
proof -
have "B.VVV.ide (DN \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> VVV_ide_char by blast
thus ?thesis
using \<mu>\<nu>\<tau> \<a>_def B.\<alpha>'.components_are_iso by force
qed
qed
sublocale bicategory V H \<a> \<i> src trg
proof
show "\<And>a. obj a \<Longrightarrow> \<guillemotleft>\<i> a : H a a \<rightarrow>\<^sub>B a\<guillemotright>"
using obj_def objE B.obj_def B.objE B.unit_in_hom by meson
show "\<And>a. obj a \<Longrightarrow> B.iso (\<i> a)"
using obj_def objE B.obj_def B.objE B.iso_unit by meson
show "\<And>f g h k. \<lbrakk> B.ide f; B.ide g; B.ide h; B.ide k;
src f = trg g; src g = trg h; src h = trg k \<rbrakk> \<Longrightarrow>
(f \<star> \<a>[g, h, k]) \<cdot> \<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k) = \<a>[f, g, h \<star> k] \<cdot> \<a>[f \<star> g, h, k]"
unfolding \<a>_def
- using B.\<a>'_def B.comp_assoc B.pentagon' VVV.arr_char VV.arr_char by simp
+ using B.\<a>'_def B.comp_assoc B.pentagon' VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
proposition is_bicategory:
shows "bicategory V H \<a> \<i> 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 "\<a>[f, g, h] = B.\<a>' h g f"
using assms \<a>_def B.\<a>'_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 "\<guillemotleft>lunit f : H (trg f) f \<rightarrow>\<^sub>B f\<guillemotright>"
using assms by simp
show "trg f \<star> lunit f = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sub>B[f, trg f, trg f]"
proof -
have "trg f \<star> lunit f = (\<i>[trg f] \<star> f) \<cdot> \<a>' (trg f) (trg f) f"
using assms lunit_char(1-2) [of f] by simp
moreover have "\<a>' (trg f) (trg f) f = \<a>\<^sub>B[f, trg f, trg f]"
proof (unfold \<a>'_def)
have 1: "VVV.ide (trg f, trg f, f)"
- using assms VVV.ide_char VVV.arr_char VV.arr_char by simp
+ using assms VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have "\<alpha>' (trg f, trg f, f) = B.inv \<a>[trg f, trg f, f]"
using 1 B.\<alpha>'.inverts_components by simp
also have "... = B.inv (B.\<alpha>' (f, trg f, trg f))"
unfolding \<a>_def using 1 by simp
also have "... = \<a>\<^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.\<alpha>'.inverts_components B.\<alpha>_def
+ using 1 B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.ide_char\<^sub>S\<^sub>b\<^sub>C
+ VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.\<alpha>'.inverts_components B.\<alpha>_def
by force
finally show "\<alpha>' (trg f, trg f, f) = \<a>\<^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.\<a>'_def assoc_ide_simp
by (intro B.lunit_eqI) auto
end
context pseudofunctor
begin
interpretation C': op_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C ..
interpretation D': op_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
notation C'.H (infixr "\<star>\<^sub>C\<^sup>o\<^sup>p" 53)
notation D'.H (infixr "\<star>\<^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
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D\<^sup>o\<^sup>p snd \<mu>\<nu>\<close> ..
interpretation FoH\<^sub>C': composite_functor C'.VV.comp V\<^sub>C V\<^sub>D
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>C\<^sup>o\<^sup>p snd \<mu>\<nu>\<close> F
..
interpretation \<Phi>': natural_isomorphism C'.VV.comp V\<^sub>D H\<^sub>D'oFF.map FoH\<^sub>C'.map
\<open>\<lambda>f. \<Phi> (snd f, fst f)\<close>
- using C.VV.arr_char C'.VV.arr_char C'.VV.ide_char C.VV.ide_char FF_def F'.FF_def
+ using C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C FF_def F'.FF_def
\<Phi>.is_extensional \<Phi>.is_natural_1 \<Phi>.is_natural_2
C.VV.dom_simp C.VV.cod_simp C'.VV.dom_simp C'.VV.cod_simp
by unfold_locales auto
interpretation F': pseudofunctor V\<^sub>C C'.H C'.\<a> \<i>\<^sub>C C'.src C'.trg
V\<^sub>D D'.H D'.\<a> \<i>\<^sub>D D'.src D'.trg
F \<open>\<lambda>f. \<Phi> (snd f, fst f)\<close>
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'.\<a> f g h) \<cdot>\<^sub>D \<Phi> (snd (f \<star>\<^sub>C\<^sup>o\<^sup>p g, h), fst (f \<star>\<^sub>C\<^sup>o\<^sup>p g, h)) \<cdot>\<^sub>D
(\<Phi> (snd (f, g), fst (f, g)) \<star>\<^sub>D\<^sup>o\<^sup>p F h) =
\<Phi> (snd (f, g \<star>\<^sub>C\<^sup>o\<^sup>p h), fst (f, g \<star>\<^sub>C\<^sup>o\<^sup>p h)) \<cdot>\<^sub>D (F f \<star>\<^sub>D\<^sup>o\<^sup>p \<Phi> (snd (g, h), fst (g, h))) \<cdot>\<^sub>D
D'.\<a> (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
+ using f g h fg gh C.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<a>\<^sub>C h g f)" "\<Phi> (g \<star>\<^sub>C\<^sup>o\<^sup>p h, f) \<cdot>\<^sub>D (F f \<star>\<^sub>D\<^sup>o\<^sup>p \<Phi> (h, g))"
"\<Phi> (h, f \<star>\<^sub>C\<^sup>o\<^sup>p g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D\<^sup>o\<^sup>p F h)" "\<a>\<^sub>D (F h) (F g) (F f)"]
D.comp_assoc
by auto
qed
lemma induces_pseudofunctor_between_opposites:
shows "pseudofunctor (\<cdot>\<^sub>C) (\<star>\<^sub>C\<^sup>o\<^sup>p) C'.\<a> \<i>\<^sub>C C'.src C'.trg
(\<cdot>\<^sub>D) (\<star>\<^sub>D\<^sup>o\<^sup>p) D'.\<a> \<i>\<^sub>D D'.src D'.trg
F (\<lambda>f. \<Phi> (snd f, fst f))"
..
text \<open>
It is now easy to dualize the coherence condition for \<open>F\<close> with respect to
left unitors to obtain the corresponding condition for right unitors.
\<close>
lemma runit_coherence:
assumes "C.ide f"
shows "\<r>\<^sub>D[F f] = F \<r>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "C'.lunit f = \<r>\<^sub>C[f]"
using assms C'.lunit_ide_simp by simp
moreover have "D'.lunit (F f) = \<r>\<^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
moreover have "D'.lunit (F f) =
F (C'.lunit f) \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^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 \<open>
The objective of this section is to establish explicit formulas for the result
of applying a pseudofunctor to expressions of various forms.
\<close>
context pseudofunctor
begin
lemma preserves_lunit:
assumes "C.ide f"
shows "F \<l>\<^sub>C[f] = \<l>\<^sub>D[F f] \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))"
and "F \<l>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]"
proof -
show 1: "F \<l>\<^sub>C[f] = \<l>\<^sub>D[F f] \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))"
proof -
have "\<l>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) = F \<l>\<^sub>C[f]"
proof -
have "D.arr \<l>\<^sub>D[F f]"
using assms by simp
moreover have "\<l>\<^sub>D[F f] = F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
using assms lunit_coherence by simp
moreover have "D.iso (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^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 (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) =
(D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))"
- using assms C.VV.arr_char unit_char FF_def D.inv_comp C.VV.dom_simp by simp
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C unit_char FF_def D.inv_comp C.VV.dom_simp by simp
ultimately show ?thesis by simp
qed
show "F \<l>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]"
proof -
have "F \<l>\<^sub>C\<^sup>-\<^sup>1[f] =
D.inv (\<l>\<^sub>D[F f] \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f)))"
using assms 1 preserves_inv by simp
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^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 \<r>\<^sub>C[f] = \<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
and "F \<r>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f]"
proof -
show 1: "F \<r>\<^sub>C[f] = \<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F \<r>\<^sub>C[f] = \<r>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)))"
proof -
have "\<r>\<^sub>D[F f] = F \<r>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
using assms runit_coherence by simp
moreover have "D.iso (\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^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 \<r>\<^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 (\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))) =
(F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (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
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C unit_char D.iso_hcomp FF_def D.inv_comp C.VV.dom_simp
by simp
ultimately show ?thesis by simp
qed
show "F \<r>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f]"
proof -
have "F \<r>\<^sub>C\<^sup>-\<^sup>1[f] =
D.inv (\<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f)))"
using assms 1 preserves_inv C.iso_runit by simp
also have "... = \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^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 \<a>\<^sub>C[f, g, h] = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
and "F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, h] = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
show 1: "F \<a>\<^sub>C[f, g, h] =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
proof -
have "F \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
using assms assoc_coherence [of f g h] by simp
moreover have "D.seq (\<Phi> (f, g \<star>\<^sub>C h)) ((F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^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
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def C.VV.dom_simp C.VV.cod_simp by auto
moreover have 2: "D.iso (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h))"
- using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def C.VV.dom_simp C.VV.cod_simp by auto
moreover have "D.inv (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)) =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
- using assms 2 C.VV.arr_char FF_def D.inv_comp C.VV.dom_simp C.VV.cod_simp
+ using assms 2 C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def D.inv_comp C.VV.dom_simp C.VV.cod_simp
by simp
ultimately show ?thesis
using D.invert_side_of_triangle(2)
[of "\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
"F \<a>\<^sub>C[f, g, h]" "\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)"]
D.comp_assoc
by simp
qed
show "F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, h] =
\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
have "F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, h] =
D.inv (\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h)))"
using assms 1 preserves_inv by simp
also have "... = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
have "\<guillemotleft>\<Phi> (f, g \<star>\<^sub>C h) : F f \<star>\<^sub>D F (g \<star>\<^sub>C h) \<Rightarrow>\<^sub>D F (f \<star>\<^sub>C g \<star>\<^sub>C h)\<guillemotright> \<and> D.iso (\<Phi> (f, g \<star>\<^sub>C h))"
using assms by auto
moreover have "\<guillemotleft>F f \<star>\<^sub>D \<Phi> (g, h) : F f \<star>\<^sub>D F g \<star>\<^sub>D F h \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F (g \<star>\<^sub>C h)\<guillemotright> \<and>
D.iso (F f \<star>\<^sub>D \<Phi> (g, h))"
using assms
by (intro conjI D.hcomp_in_vhom, auto)
ultimately show ?thesis
using assms D.isos_compose D.comp_assoc D.inv_comp
by (elim conjE D.in_homE) auto
qed
finally show ?thesis by simp
qed
qed
lemma preserves_hcomp:
assumes "C.hseq \<mu> \<nu>"
shows "F (\<mu> \<star>\<^sub>C \<nu>) =
\<Phi> (C.cod \<mu>, C.cod \<nu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D F \<nu>) \<cdot>\<^sub>D D.inv (\<Phi> (C.dom \<mu>, C.dom \<nu>))"
proof -
have "F (\<mu> \<star>\<^sub>C \<nu>) \<cdot>\<^sub>D \<Phi> (C.dom \<mu>, C.dom \<nu>) = \<Phi> (C.cod \<mu>, C.cod \<nu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D F \<nu>)"
- using assms \<Phi>.naturality C.VV.arr_char C.VV.cod_char FF_def C.VV.dom_simp
+ using assms \<Phi>.naturality C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C 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 \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
shows "adjunction_data_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F f) (F g) (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof
interpret adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
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 "\<guillemotleft>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) : src\<^sub>D (F f) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F f\<guillemotright>"
- using antipar C.VV.ide_char C.VV.arr_char cmp_in_hom(2) unit_in_hom FF_def by auto
+ using antipar C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_in_hom(2) unit_in_hom FF_def by auto
show "\<guillemotleft>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) : F f \<star>\<^sub>D F g \<Rightarrow>\<^sub>D src\<^sub>D (F g)\<guillemotright>"
- using antipar C.VV.ide_char C.VV.arr_char FF_def cmp_in_hom(2) unit_in_hom(2)
+ using antipar C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
shows "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F f) (F g) (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof -
interpret equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms by auto
show "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F f) (F g) (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
- using antipar unit_is_iso preserves_iso unit_char(2) C.VV.ide_char C.VV.arr_char
+ using antipar unit_is_iso preserves_iso unit_char(2) C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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 \<eta> \<epsilon> where E: "equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
using assms C.equivalence_map_def by auto
interpret E: equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
begin
text\<open>
The underlying vertical functor is just the identity functor on the vertical category,
which is already available as \<open>B.map\<close>.
\<close>
abbreviation map
where "map \<equiv> 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
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
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close>
..
interpretation IoH\<^sub>B: composite_functor B.VV.comp V\<^sub>B V\<^sub>B \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close> map
..
text\<open>
The horizontal composition provides the compositor.
\<close>
abbreviation cmp
where "cmp \<equiv> \<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>"
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
+ using B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^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 \<and> 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 \<open>
In this section, we construct the embedding pseudofunctor of a sub-bicategory
into the ambient bicategory.
\<close>
locale embedding_pseudofunctor =
B: bicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B +
S: subbicategory
begin
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
definition map
where "map \<mu> = (if S.arr \<mu> then \<mu> else B.null)"
lemma map_in_hom [intro]:
assumes "S.arr \<mu>"
shows "\<guillemotleft>map \<mu> : src\<^sub>B (map (S.src \<mu>)) \<rightarrow>\<^sub>B src\<^sub>B (map (S.trg \<mu>))\<guillemotright>"
and "\<guillemotleft>map \<mu> : map (S.dom \<mu>) \<Rightarrow>\<^sub>B map (S.cod \<mu>)\<guillemotright>"
proof -
show 1: "\<guillemotleft>map \<mu> : map (S.dom \<mu>) \<Rightarrow>\<^sub>B map (S.cod \<mu>)\<guillemotright>"
- using assms map_def S.in_hom_char by fastforce
+ using assms map_def S.in_hom_char\<^sub>S\<^sub>b\<^sub>C by fastforce
show "\<guillemotleft>map \<mu> : src\<^sub>B (map (S.src \<mu>)) \<rightarrow>\<^sub>B src\<^sub>B (map (S.trg \<mu>))\<guillemotright>"
- using assms 1 map_def S.arr_char S.src_def S.trg_def S.obj_char S.obj_src S.obj_trg
+ using assms 1 map_def S.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu>"
shows "B.arr (map \<mu>)"
and "src\<^sub>B (map \<mu>) = src\<^sub>B (map (S.src \<mu>))" and "trg\<^sub>B (map \<mu>) = src\<^sub>B (map (S.trg \<mu>))"
and "B.dom (map \<mu>) = map (S.dom \<mu>)" and "B.cod (map \<mu>) = map (S.cod \<mu>)"
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
+ using map_def S.comp_char S.seq_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C
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
+ using S.arr_char\<^sub>S\<^sub>b\<^sub>C 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
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close>
..
interpretation FoH: composite_functor S.VV.comp S.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> map
..
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
definition cmp
where "cmp \<mu>\<nu> = (if S.VV.arr \<mu>\<nu> then fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu> else B.null)"
lemma cmp_in_hom [intro]:
assumes "S.VV.arr \<mu>\<nu>"
shows "\<guillemotleft>cmp \<mu>\<nu> : src\<^sub>B (snd \<mu>\<nu>) \<rightarrow>\<^sub>B trg\<^sub>B (fst \<mu>\<nu>)\<guillemotright>"
and "\<guillemotleft>cmp \<mu>\<nu> : map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))\<guillemotright>"
proof -
show "\<guillemotleft>cmp \<mu>\<nu> : map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))\<guillemotright>"
proof
show 1: "B.arr (cmp \<mu>\<nu>)"
unfolding cmp_def
- using assms S.arr_char S.VV.arr_char S.inclusion S.src_def S.trg_def by auto
+ using assms S.arr_char\<^sub>S\<^sub>b\<^sub>C S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.inclusion S.src_def S.trg_def by auto
show "B.dom (cmp \<mu>\<nu>) = map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))"
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
+ using assms 1 cmp_def S.dom_simp S.cod_simp S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C S.hcomp_def
S.inclusion S.dom_closed
by auto
show "B.cod (cmp \<mu>\<nu>) = map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))"
unfolding cmp_def map_def
using assms 1 S.H.preserves_arr S.cod_simp S.hcomp_eqI S.hcomp_simps(4) S.hseq_char'
by auto
qed
thus "\<guillemotleft>cmp \<mu>\<nu> : src\<^sub>B (snd \<mu>\<nu>) \<rightarrow>\<^sub>B trg\<^sub>B (fst \<mu>\<nu>)\<guillemotright>"
using cmp_def by auto
qed
lemma cmp_simps [simp]:
assumes "S.VV.arr \<mu>\<nu>"
shows "B.arr (cmp \<mu>\<nu>)"
and "src\<^sub>B (cmp \<mu>\<nu>) = S.src (snd \<mu>\<nu>)" and "trg\<^sub>B (cmp \<mu>\<nu>) = S.trg (fst \<mu>\<nu>)"
and "B.dom (cmp \<mu>\<nu>) = map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))"
and "B.cod (cmp \<mu>\<nu>) = map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))"
- using assms cmp_in_hom S.src_def S.trg_def S.VV.arr_char
+ using assms cmp_in_hom S.src_def S.trg_def S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp_all blast+
lemma iso_cmp:
assumes "S.VV.ide \<mu>\<nu>"
shows "B.iso (cmp \<mu>\<nu>)"
- using assms S.VV.ide_char S.VV.arr_char S.arr_char cmp_def S.ide_char S.src_def S.trg_def
+ using assms S.VV.ide_char\<^sub>S\<^sub>b\<^sub>C S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_def S.ide_char\<^sub>S\<^sub>b\<^sub>C S.src_def S.trg_def
by auto
interpretation \<Phi>\<^sub>E: natural_isomorphism S.VV.comp V HoFF.map FoH.map cmp
proof
show "\<And>\<mu>\<nu>. \<not> S.VV.arr \<mu>\<nu> \<Longrightarrow> cmp \<mu>\<nu> = B.null"
using cmp_def by simp
fix \<mu>\<nu>
assume \<mu>\<nu>: "S.VV.arr \<mu>\<nu>"
have 1: "S.arr (fst \<mu>\<nu>) \<and> S.arr (snd \<mu>\<nu>) \<and> S.src (fst \<mu>\<nu>) = S.trg (snd \<mu>\<nu>)"
- using \<mu>\<nu> S.VV.arr_char by simp
+ using \<mu>\<nu> S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show "B.dom (cmp \<mu>\<nu>) = HoFF.map (S.VV.dom \<mu>\<nu>)"
- using \<mu>\<nu> 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
+ using \<mu>\<nu> FF_def S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.VV.dom_char\<^sub>S\<^sub>b\<^sub>C S.arr_dom S.src_def S.trg_def
+ S.dom_char\<^sub>S\<^sub>b\<^sub>C S.src.preserves_dom S.trg.preserves_dom
apply simp
by (metis (no_types, lifting))
show "B.cod (cmp \<mu>\<nu>) = FoH.map (S.VV.cod \<mu>\<nu>)"
- using \<mu>\<nu> 1 map_def S.hseq_char S.hcomp_def S.cod_char S.arr_cod S.VV.cod_simp
+ using \<mu>\<nu> 1 map_def S.hseq_char S.hcomp_def S.cod_char\<^sub>S\<^sub>b\<^sub>C S.arr_cod S.VV.cod_simp
by simp
show "cmp (S.VV.cod \<mu>\<nu>) \<cdot>\<^sub>B HoFF.map \<mu>\<nu> = cmp \<mu>\<nu>"
- using \<mu>\<nu> 1 cmp_def S.VV.arr_char S.VV.cod_char FF_def S.arr_cod S.cod_simp
+ using \<mu>\<nu> 1 cmp_def S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.VV.cod_char\<^sub>S\<^sub>b\<^sub>C 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) \<mu>\<nu>)
show "FoH.map \<mu>\<nu> \<cdot>\<^sub>B cmp (S.VV.dom \<mu>\<nu>) = cmp \<mu>\<nu>"
unfolding cmp_def map_def
- using \<mu>\<nu> S.VV.arr_char B.VV.arr_char S.VV.dom_char S.VV.cod_char B.comp_arr_dom
+ using \<mu>\<nu> S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.VV.dom_char\<^sub>S\<^sub>b\<^sub>C S.VV.cod_char\<^sub>S\<^sub>b\<^sub>C 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)
+ by (metis (no_types, lifting) B.hcomp_simps(3) cmp_def cmp_simps(1) S.arr_char\<^sub>S\<^sub>b\<^sub>C
+ S.dom_char\<^sub>S\<^sub>b\<^sub>C S.hcomp_closed S.src_def S.trg_def)
next
show "\<And>fg. S.VV.ide fg \<Longrightarrow> B.iso (cmp fg)"
using iso_cmp by simp
qed
sublocale pseudofunctor S.comp S.hcomp S.\<a> \<i> S.src S.trg V H \<a>\<^sub>B \<i> 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 \<and> B.ide g \<and> B.ide h \<and> src\<^sub>B f = trg\<^sub>B g \<and> src\<^sub>B g = trg\<^sub>B h"
- using f g h fg gh S.ide_char S.arr_char S.src_def S.trg_def by simp
+ using f g h fg gh S.ide_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C S.src_def S.trg_def by simp
show "map (S.\<a> f g h) \<cdot>\<^sub>B cmp (f \<star> g, h) \<cdot>\<^sub>B cmp (f, g) \<star>\<^sub>B map h =
cmp (f, g \<star> h) \<cdot>\<^sub>B (map f \<star>\<^sub>B cmp (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[map f, map g, map h]"
proof -
have "map (S.\<a> f g h) \<cdot>\<^sub>B cmp (f \<star> g, h) \<cdot>\<^sub>B cmp (f, g) \<star>\<^sub>B map h =
\<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^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
+ using 1 f g h fg gh S.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.comp_arr_dom S.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<star> h) \<cdot>\<^sub>B (map f \<star>\<^sub>B cmp (g, h)) \<cdot>\<^sub>B \<a>\<^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
+ using 1 f g h fg gh S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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
+ S.arr_char\<^sub>S\<^sub>b\<^sub>C 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.\<a> \<i> S.src S.trg V H \<a>\<^sub>B \<i> 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 "\<guillemotleft>a : map\<^sub>0 a \<Rightarrow>\<^sub>B map a\<guillemotright>"
- using assms a 2 map\<^sub>0_def map_def S.src_def S.dom_char S.cod_char S.obj_char
+ using assms a 2 map\<^sub>0_def map_def S.src_def S.dom_char\<^sub>S\<^sub>b\<^sub>C S.cod_char\<^sub>S\<^sub>b\<^sub>C S.obj_char
by auto
show "B.iso a"
using assms 1 by simp
show "a \<cdot>\<^sub>B \<i>[map\<^sub>0 a] = (map \<i>[a] \<cdot>\<^sub>B cmp (a, a)) \<cdot>\<^sub>B (a \<star>\<^sub>B a)"
proof -
have "a \<cdot>\<^sub>B \<i>[a] = \<i>[a] \<cdot>\<^sub>B cmp (a, a) \<cdot>\<^sub>B (a \<star>\<^sub>B a)"
proof -
have "a \<cdot>\<^sub>B \<i>[a] = \<i>[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 \<star>\<^sub>B a) \<cdot>\<^sub>B (a \<star>\<^sub>B a) = a \<star>\<^sub>B a"
using assms S.obj_char S.comp_ide_self by auto
moreover have "B.dom \<i>[a] = a \<star>\<^sub>B a"
using assms S.obj_char by simp
moreover have "\<i>[a] \<cdot>\<^sub>B (a \<star>\<^sub>B a) = \<i>[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
+ using assms cmp_def S.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<open>
In this section, we show how pseudofunctors may be composed. The main work is to
establish the coherence condition for associativity.
\<close>
locale composite_pseudofunctor =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F +
G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
and V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'b \<Rightarrow> 'c"
and \<Phi>\<^sub>F :: "'b * 'b \<Rightarrow> 'c"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> '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 \<open>G o F\<close>
proof
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> D.isomorphic ((G o F) (src\<^sub>B \<mu>)) (src\<^sub>D ((G o F) \<mu>))"
proof -
fix \<mu>
assume \<mu>: "B.arr \<mu>"
show "D.isomorphic ((G o F) (src\<^sub>B \<mu>)) (src\<^sub>D ((G o F) \<mu>))"
proof -
have "(G o F) (src\<^sub>B \<mu>) = G (F (src\<^sub>B \<mu>))"
using \<mu> by simp
also have "D.isomorphic ... (G (src\<^sub>C (F \<mu>)))"
using \<mu> 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 \<mu>)))"
using \<mu> G.weakly_preserves_src by blast
also have "... = src\<^sub>D ((G o F) \<mu>)"
by simp
finally show ?thesis by blast
qed
qed
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> D.isomorphic ((G o F) (trg\<^sub>B \<mu>)) (trg\<^sub>D ((G o F) \<mu>))"
proof -
fix \<mu>
assume \<mu>: "B.arr \<mu>"
show "D.isomorphic ((G o F) (trg\<^sub>B \<mu>)) (trg\<^sub>D ((G o F) \<mu>))"
proof -
have "(G o F) (trg\<^sub>B \<mu>) = G (F (trg\<^sub>B \<mu>))"
using \<mu> by simp
also have "D.isomorphic ... (G (trg\<^sub>C (F \<mu>)))"
using \<mu> 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 \<mu>)))"
using \<mu> G.weakly_preserves_trg by blast
also have "... = trg\<^sub>D ((G o F) \<mu>)"
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
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D snd \<mu>\<nu>\<close>
..
interpretation GFoH\<^sub>B: composite_functor B.VV.comp V\<^sub>B V\<^sub>D \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close>
\<open>G o F\<close>
..
definition cmp
where "cmp \<mu>\<nu> = (if B.VV.arr \<mu>\<nu> then
G (F (H\<^sub>B (fst \<mu>\<nu>) (snd \<mu>\<nu>))) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.dom (fst \<mu>\<nu>)), F (B.dom (snd \<mu>\<nu>)))
else D.null)"
lemma cmp_in_hom [intro]:
assumes "B.VV.arr \<mu>\<nu>"
shows "\<guillemotleft>cmp \<mu>\<nu> : H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>) \<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)\<guillemotright>"
proof -
have "cmp \<mu>\<nu> = G (F (H\<^sub>B (fst \<mu>\<nu>) (snd \<mu>\<nu>))) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.dom (fst \<mu>\<nu>)), F (B.dom (snd \<mu>\<nu>)))"
using assms cmp_def by simp
moreover have "\<guillemotleft> ... : H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>) \<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<Phi>\<^sub>G (F (B.dom (fst \<mu>\<nu>)), F (B.dom (snd \<mu>\<nu>))) :
H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)
\<Rightarrow>\<^sub>D G (F (B.dom (fst \<mu>\<nu>)) \<star>\<^sub>C F (B.dom (snd \<mu>\<nu>)))\<guillemotright>"
- using assms F.FF_def FF_def B.VV.arr_char B.VV.dom_simp by auto
+ using assms F.FF_def FF_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp by auto
show "\<guillemotleft>G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) :
G (F (B.dom (fst \<mu>\<nu>)) \<star>\<^sub>C F (B.dom (snd \<mu>\<nu>)))
\<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.dom \<mu>\<nu>)\<guillemotright>"
- using assms B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
+ using assms B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
show "\<guillemotleft>G (F (fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)) :
GFoH\<^sub>B.map (B.VV.dom \<mu>\<nu>) \<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)\<guillemotright>"
proof -
have "B.VV.in_hom \<mu>\<nu> (B.VV.dom \<mu>\<nu>) (B.VV.cod \<mu>\<nu>)"
using assms by auto
thus ?thesis by auto
qed
qed
ultimately show ?thesis by argo
qed
lemma cmp_simps [simp]:
assumes "B.VV.arr \<mu>\<nu>"
shows "D.arr (cmp \<mu>\<nu>)"
and "D.dom (cmp \<mu>\<nu>) = H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)"
and "D.cod (cmp \<mu>\<nu>) = GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)"
using assms cmp_in_hom by blast+
interpretation \<Phi>: natural_transformation
B.VV.comp V\<^sub>D H\<^sub>DoGF_GF.map GFoH\<^sub>B.map cmp
proof
show "\<And>\<mu>\<nu>. \<not> B.VV.arr \<mu>\<nu> \<Longrightarrow> cmp \<mu>\<nu> = D.null"
unfolding cmp_def by simp
fix \<mu>\<nu>
assume \<mu>\<nu>: "B.VV.arr \<mu>\<nu>"
show "D.dom (cmp \<mu>\<nu>) = H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)"
using \<mu>\<nu> cmp_in_hom by blast
show "D.cod (cmp \<mu>\<nu>) = GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)"
using \<mu>\<nu> cmp_in_hom by blast
show "GFoH\<^sub>B.map \<mu>\<nu> \<cdot>\<^sub>D cmp (B.VV.dom \<mu>\<nu>) = cmp \<mu>\<nu>"
unfolding cmp_def
- using \<mu>\<nu> B.VV.ide_char B.VV.arr_char D.comp_ide_arr B.VV.dom_char D.comp_assoc
+ using \<mu>\<nu> B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.comp_ide_arr B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu>\<nu>) \<cdot>\<^sub>D H\<^sub>DoGF_GF.map \<mu>\<nu> = cmp \<mu>\<nu>"
proof -
have "cmp (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>D H\<^sub>DoGF_GF.map \<mu>\<nu> =
(G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>)))) \<cdot>\<^sub>D
(fst (FF \<mu>\<nu>) \<star>\<^sub>D snd (FF \<mu>\<nu>))"
unfolding cmp_def
- using \<mu>\<nu> B.VV.arr_char B.VV.dom_simp B.VV.cod_simp by simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp by simp
also have "... = (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>)))) \<cdot>\<^sub>D
(fst (FF \<mu>\<nu>) \<star>\<^sub>D snd (FF \<mu>\<nu>))"
proof -
have "D.ide (G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))))"
- using \<mu>\<nu> B.VV.arr_char by simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "D.seq (G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))))
(G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))))"
- using \<mu>\<nu> B.VV.arr_char B.VV.dom_char B.VV.cod_char F.FF_def
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
apply (intro D.seqI)
apply auto
proof -
show "G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))) =
D.cod (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))))"
proof -
have "D.seq (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))))
(\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))))"
- using \<mu>\<nu> B.VV.arr_char F.FF_def B.VV.arr_char B.VV.dom_char B.VV.cod_char
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by (intro D.seqI) auto
thus ?thesis
- using \<mu>\<nu> B.VV.arr_char B.VV.cod_simp by simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_simp by simp
qed
qed
ultimately show ?thesis
using \<mu>\<nu> D.comp_ide_arr [of "G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)))"
"G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>)))"]
by simp
qed
also have "... = G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
(fst (FF \<mu>\<nu>) \<star>\<^sub>D snd (FF \<mu>\<nu>)))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.cod (F.FF \<mu>\<nu>)) \<cdot>\<^sub>D G.H\<^sub>DoFF.map (F.FF \<mu>\<nu>)"
- using \<mu>\<nu> B.VV.arr_char F.FF_def G.FF_def FF_def C.VV.cod_simp by auto
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def G.FF_def FF_def C.VV.cod_simp by auto
also have "... = G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
G.FoH\<^sub>C.map (F.FF \<mu>\<nu>) \<cdot>\<^sub>D \<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
- using \<mu>\<nu> B.VV.arr_char G.\<Phi>.naturality C.VV.dom_simp C.VV.cod_simp by simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C G.\<Phi>.naturality C.VV.dom_simp C.VV.cod_simp by simp
also have "... = (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
G.FoH\<^sub>C.map (F.FF \<mu>\<nu>)) \<cdot>\<^sub>D \<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
using D.comp_assoc by simp
also have "... = (G (\<Phi>\<^sub>F (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>C F.H\<^sub>DoFF.map \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
proof -
have "(B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>)) = B.VV.cod \<mu>\<nu>"
- using \<mu>\<nu> B.VV.arr_char B.VV.cod_simp by simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_simp by simp
moreover have "G.FoH\<^sub>C.map (F.FF \<mu>\<nu>) = G (F.H\<^sub>DoFF.map \<mu>\<nu>)"
using \<mu>\<nu> F.FF_def by simp
moreover have "G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D G (F.H\<^sub>DoFF.map \<mu>\<nu>) =
G (\<Phi>\<^sub>F (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>C F.H\<^sub>DoFF.map \<mu>\<nu>)"
- using \<mu>\<nu> B.VV.arr_char
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) F.\<Phi>.is_natural_2 F.\<Phi>.preserves_reflects_arr
G.preserves_comp calculation(1))
ultimately show ?thesis by argo
qed
also have "... = G (F.FoH\<^sub>C.map \<mu>\<nu> \<cdot>\<^sub>C \<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
using \<mu>\<nu> F.\<Phi>.naturality [of \<mu>\<nu>] F.FF_def by simp
also have "... = G (F.FoH\<^sub>C.map \<mu>\<nu>) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
proof -
have "G (F.FoH\<^sub>C.map \<mu>\<nu> \<cdot>\<^sub>C \<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) =
G (F.FoH\<^sub>C.map \<mu>\<nu>) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>))"
using \<mu>\<nu>
by (metis (mono_tags, lifting) F.\<Phi>.is_natural_1 F.\<Phi>.preserves_reflects_arr
G.preserves_comp)
thus ?thesis
using \<mu>\<nu> D.comp_assoc by simp
qed
also have "... = cmp \<mu>\<nu>"
- using \<mu>\<nu> B.VV.arr_char cmp_def F.FF_def F.FF.preserves_dom B.VV.dom_simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C cmp_def F.FF_def F.FF.preserves_dom B.VV.dom_simp
by auto
finally show ?thesis by simp
qed
qed
interpretation \<Phi>: natural_isomorphism B.VV.comp V\<^sub>D H\<^sub>DoGF_GF.map GFoH\<^sub>B.map cmp
proof
show "\<And>hk. B.VV.ide hk \<Longrightarrow> D.iso (cmp hk)"
proof -
fix hk
assume hk: "B.VV.ide hk"
have "D.iso (G (F (fst hk \<star>\<^sub>B snd hk)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom hk)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
proof (intro D.isos_compose)
show "D.iso (\<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
using hk G.\<Phi>.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
+ C.VV.ide_char B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C
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.\<Phi>.preserves_iso fst_conv snd_conv)
show "D.iso (G (\<Phi>\<^sub>F (B.VV.dom hk)))"
- using hk F.\<Phi>.components_are_iso B.VV.arr_char B.VV.dom_char B.VV.ideD(2)
+ using hk F.\<Phi>.components_are_iso B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.ideD(2)
by auto
show "D.seq (G (\<Phi>\<^sub>F (B.VV.dom hk))) (\<Phi>\<^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
+ using hk B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
C.VV.dom_simp C.VV.cod_simp
by auto
show "D.iso (G (F (fst hk \<star>\<^sub>B snd hk)))"
- using hk F.\<Phi>.components_are_iso B.VV.arr_char by simp
+ using hk F.\<Phi>.components_are_iso B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show "D.seq (G (F (fst hk \<star>\<^sub>B snd hk)))
(G (\<Phi>\<^sub>F (B.VV.dom hk)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
- using hk B.VV.arr_char B.VV.dom_char
+ using hk B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C
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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>G o F\<close> 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 \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>D cmp (f \<star>\<^sub>B g, h) \<cdot>\<^sub>D (cmp (f, g) \<star>\<^sub>D map h) =
cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (map f \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[map f, map g, map h]"
proof -
have "map \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>D cmp (f \<star>\<^sub>B g, h) \<cdot>\<^sub>D (cmp (f, g) \<star>\<^sub>D map h) =
G (F \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D
(G (F ((f \<star>\<^sub>B g) \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h)) \<cdot>\<^sub>D
(G (F (f \<star>\<^sub>B g)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
unfolding cmp_def
- using f g h fg gh B.VV.arr_char B.VV.dom_simp by simp
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp by simp
also have "... = G (F \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h)) \<cdot>\<^sub>D
(G (F (f \<star>\<^sub>B g)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^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 \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h)) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^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 \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
- using f g h fg gh B.VV.arr_char B.VV.cod_simp by simp
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_simp by simp
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<star>\<^sub>D G (F h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def D.whisker_right
B.VV.dom_simp C.VV.cod_simp
by auto
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<star>\<^sub>D G (F h))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
proof -
have "\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) = \<Phi>\<^sub>G (C.VV.cod (\<Phi>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_simp C.VV.cod_simp
by simp
moreover have "G (\<Phi>\<^sub>F (f, g)) \<star>\<^sub>D G (F h) = G.H\<^sub>DoFF.map (\<Phi>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C G.FF_def by simp
moreover have "G.FoH\<^sub>C.map (\<Phi>\<^sub>F (f, g), F h) = G (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h)"
- using f g h fg gh B.VV.arr_char by simp
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "\<Phi>\<^sub>G (C.VV.dom (\<Phi>\<^sub>F (f, g), F h)) = \<Phi>\<^sub>G (F f \<star>\<^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
+ using f g h fg gh C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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.\<Phi>.naturality
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C G.\<Phi>.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.\<Phi>.is_extensional)
qed
also have "... = (G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G ((F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>C (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using C.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>C \<a>\<^sub>C[F f, F g, F h]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh F.assoc_coherence [of f g h] by simp
also have "... = G ((\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>C \<a>\<^sub>C[F f, F g, F h]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using C.comp_assoc by simp
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D G \<a>\<^sub>C[F f, F g, F h]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
G \<a>\<^sub>C[F f, F g, F h] \<cdot>\<^sub>D \<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h)) \<cdot>\<^sub>D
\<a>\<^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 (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F f), G (F g), G (F h)]"
using D.comp_assoc by simp
also have "... = (cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h))) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F f), G (F g), G (F h)]"
proof -
have "G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h)) =
cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h))"
proof -
have "cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h)) =
(G (F (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h))) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp by simp
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h))) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
proof -
have "G (F (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) = G (\<Phi>\<^sub>F (f, g \<star>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
proof -
have "G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) = G (\<Phi>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using f g h fg gh
D.whisker_left [of "G (F f)" "G (\<Phi>\<^sub>F (g, h))" "\<Phi>\<^sub>G (F g, F h)"]
by fastforce
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h)))) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
proof -
have "\<Phi>\<^sub>G (C.VV.cod (F f, \<Phi>\<^sub>F (g, h))) = \<Phi>\<^sub>G (F f, F (g \<star>\<^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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_simp B.VV.cod_simp
by auto
moreover have "G.H\<^sub>DoFF.map (F f, \<Phi>\<^sub>F (g, h)) = G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h))"
- using f g h fg gh B.VV.arr_char G.FF_def by auto
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C G.FF_def by auto
moreover have "G.FoH\<^sub>C.map (F f, \<Phi>\<^sub>F (g, h)) = G (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))"
- using f g h fg gh B.VV.arr_char C.VV.arr_char by simp
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "C.VV.dom (F f, \<Phi>\<^sub>F (g, h)) = (F f, F g \<star>\<^sub>C F h)"
- using f g h fg gh B.VV.arr_char C.VV.arr_char C.VV.dom_char
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C
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
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
G.\<Phi>.naturality [of "(F f, \<Phi>\<^sub>F (g, h))"]
by simp
qed
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D (G (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h)))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
- using f g h fg gh B.VV.arr_char
+ using f g h fg gh B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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 \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D
\<a>\<^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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^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) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)"
proof -
have "G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) = unit a"
proof (intro unit_eqI [of a "G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)"])
show "B.obj a" by fact
show "\<guillemotleft>G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) : map\<^sub>0 a \<Rightarrow>\<^sub>D map a\<guillemotright>"
using assms by auto
show "D.iso (G (F.unit a) \<cdot>\<^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) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] =
(map \<i>\<^sub>B[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D
(G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a))"
proof -
have 1: "cmp (a, a) = G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F a, F a)"
proof -
have "cmp (a, a) = (G (F (a \<star>\<^sub>B a)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a))) \<cdot>\<^sub>D \<Phi>\<^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
+ using assms cmp_def B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
B.VxV.dom_char B.objE D.comp_assoc B.obj_simps
by simp
also have "... = G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F a, F a)"
- using assms D.comp_cod_arr B.VV.arr_char B.VV.ide_char by auto
+ using assms D.comp_cod_arr B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C by auto
finally show ?thesis by blast
qed
have "(map \<i>\<^sub>B[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D
(G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) =
map \<i>\<^sub>B[a] \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F a, F a) \<cdot>\<^sub>D (G (F.unit a) \<star>\<^sub>D G (F.unit a))) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms 1 D.comp_assoc D.interchange by simp
also have "... = (map \<i>\<^sub>B[a] \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D G (F.unit a \<star>\<^sub>C F.unit a)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms G.\<Phi>.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
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C G.FF_def D.comp_assoc
by simp
also have "... = (G (F \<i>\<^sub>B[a] \<cdot>\<^sub>C \<Phi>\<^sub>F (a, a) \<cdot>\<^sub>C (F.unit a \<star>\<^sub>C F.unit a))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
proof -
have "C.arr (F \<i>\<^sub>B[a] \<cdot>\<^sub>C \<Phi>\<^sub>F (a, a) \<cdot>\<^sub>C (F.unit a \<star>\<^sub>C F.unit a))"
- using assms B.VV.ide_char B.VV.arr_char F.cmp_in_hom(2)
+ using assms B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.cmp_in_hom(2)
by (intro C.seqI' C.comp_in_homI) auto
hence "map \<i>\<^sub>B[a] \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D G (F.unit a \<star>\<^sub>C F.unit a) =
G (F \<i>\<^sub>B[a] \<cdot>\<^sub>C \<Phi>\<^sub>F (a, a) \<cdot>\<^sub>C (F.unit a \<star>\<^sub>C F.unit a))"
by auto
thus ?thesis by argo
qed
also have "... = G (F.unit a \<cdot>\<^sub>C \<i>\<^sub>C[F.map\<^sub>0 a]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^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) \<cdot>\<^sub>D (G \<i>\<^sub>C[F.map\<^sub>0 a] \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a)) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms D.comp_assoc by simp
also have "... = (G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) \<cdot>\<^sub>D \<i>\<^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) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) \<cdot>\<^sub>D \<i>\<^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 \<open>
In this section, we construct the restriction and corestriction of a pseudofunctor to
a subbicategory of its domain and codomain, respectively.
\<close>
locale restricted_pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi> +
C': subbicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C Arr
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi> :: "'c * 'c \<Rightarrow> 'd"
and Arr :: "'c \<Rightarrow> bool"
begin
abbreviation map
where "map \<equiv> \<lambda>\<mu>. if C'.arr \<mu> then F \<mu> else D.null"
abbreviation cmp
where "cmp \<equiv> \<lambda>\<mu>\<nu>. if C'.VV.arr \<mu>\<nu> then \<Phi> \<mu>\<nu> 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
+ using C'.inclusion C'.arr_char\<^sub>S\<^sub>b\<^sub>C C'.dom_char\<^sub>S\<^sub>b\<^sub>C C'.cod_char\<^sub>S\<^sub>b\<^sub>C C'.seq_char\<^sub>S\<^sub>b\<^sub>C 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
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D snd \<mu>\<nu>\<close>
..
interpretation FoH\<^sub>C\<^sub>': composite_functor C'.VV.comp C'.comp V\<^sub>D
\<open>\<lambda>\<mu>\<nu>. C'.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> map
..
interpretation \<Phi>: 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
+ using C'.arr_char\<^sub>S\<^sub>b\<^sub>C C'.dom_char\<^sub>S\<^sub>b\<^sub>C C'.cod_char\<^sub>S\<^sub>b\<^sub>C C'.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C'.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
FF_def C'.inclusion C'.dom_closed C'.cod_closed C'.src_def C'.trg_def
C'.hcomp_def C'.hcomp_closed F.\<Phi>.is_natural_1 F.\<Phi>.is_natural_2
- C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
by unfold_locales auto
interpretation \<Phi>: 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.\<Phi>.components_are_iso
+ using C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C'.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.arr_char\<^sub>S\<^sub>b\<^sub>C
+ C'.src_def C'.trg_def C'.ide_char\<^sub>S\<^sub>b\<^sub>C F.\<Phi>.components_are_iso
by unfold_locales auto
sublocale pseudofunctor C'.comp C'.hcomp C'.\<a> \<i>\<^sub>C C'.src C'.trg V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^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
+ using F.assoc_coherence C'.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C'.arr_char\<^sub>S\<^sub>b\<^sub>C C'.hcomp_def
+ C'.src_def C'.trg_def C'.assoc_closed C'.hcomp_closed C'.ide_char\<^sub>S\<^sub>b\<^sub>C
by unfold_locales auto
lemma is_pseudofunctor:
shows "pseudofunctor C'.comp C'.hcomp C'.\<a> \<i>\<^sub>C C'.src C'.trg V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^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"
using assms map\<^sub>0_simp C'.obj_char F.unit_in_hom(2) [of a] F.unit_char(2-3) \<i>_simps(1)
apply (intro unit_eqI)
apply auto
by blast
end
text \<open>
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.
\<close>
locale corestricted_pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi> +
D': subbicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<lambda>\<mu>. D.arr \<mu> \<and> Obj (src\<^sub>D \<mu>) \<and> Obj (trg\<^sub>D \<mu>)\<close>
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi> :: "'c * 'c \<Rightarrow> 'd"
and Obj :: "'d \<Rightarrow> bool" +
assumes preserves_arr: "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> D'.arr (F \<mu>)"
begin
abbreviation map
where "map \<equiv> F"
abbreviation cmp
where "cmp \<equiv> \<Phi>"
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
+ using preserves_arr F.is_extensional D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.dom_char\<^sub>S\<^sub>b\<^sub>C D'.cod_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu>
assume \<mu>: "C.arr \<mu>"
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : F (src\<^sub>C \<mu>) \<Rightarrow>\<^sub>D src\<^sub>D (F \<mu>)\<guillemotright> \<and> D.iso \<phi>"
using \<mu> F.weakly_preserves_src by auto
have 2: "D'.in_hom \<phi> (F (src\<^sub>C \<mu>)) (D'.src (F \<mu>))"
- using \<mu> \<phi> D'.arr_char D'.dom_char D'.cod_char D'.src_def D.vconn_implies_hpar(1-2)
+ using \<mu> \<phi> D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.dom_char\<^sub>S\<^sub>b\<^sub>C D'.cod_char\<^sub>S\<^sub>b\<^sub>C D'.src_def D.vconn_implies_hpar(1-2)
preserves_arr
- by (metis (no_types, lifting) C.src.preserves_arr D'.in_hom_char D'.src.preserves_arr
+ by (metis (no_types, lifting) C.src.preserves_arr D'.in_hom_char\<^sub>S\<^sub>b\<^sub>C D'.src.preserves_arr
D.arrI)
moreover have "D'.iso \<phi>"
- using 2 \<phi> D'.iso_char D'.arr_char by fastforce
+ using 2 \<phi> D'.iso_char\<^sub>S\<^sub>b\<^sub>C D'.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
ultimately show "D'.isomorphic (F (src\<^sub>C \<mu>)) (D'.src (F \<mu>))"
using D'.isomorphic_def by auto
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : F (trg\<^sub>C \<mu>) \<Rightarrow>\<^sub>D trg\<^sub>D (F \<mu>)\<guillemotright> \<and> D.iso \<psi>"
using \<mu> F.weakly_preserves_trg by auto
have 2: "D'.in_hom \<psi> (F (trg\<^sub>C \<mu>)) (D'.trg (F \<mu>))"
- using \<mu> \<psi> D'.arr_char D'.dom_char D'.cod_char D'.trg_def D.vconn_implies_hpar(1-2)
+ using \<mu> \<psi> D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.dom_char\<^sub>S\<^sub>b\<^sub>C D'.cod_char\<^sub>S\<^sub>b\<^sub>C D'.trg_def D.vconn_implies_hpar(1-2)
preserves_arr
- by (metis (no_types, lifting) C.trg.preserves_arr D'.in_hom_char D'.trg.preserves_arr
+ by (metis (no_types, lifting) C.trg.preserves_arr D'.in_hom_char\<^sub>S\<^sub>b\<^sub>C D'.trg.preserves_arr
D.arrI)
moreover have "D'.iso \<psi>"
- using 2 \<psi> D'.iso_char D'.arr_char by fastforce
+ using 2 \<psi> D'.iso_char\<^sub>S\<^sub>b\<^sub>C D'.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
ultimately show "D'.isomorphic (F (trg\<^sub>C \<mu>)) (D'.trg (F \<mu>))"
using D'.isomorphic_def by auto
qed
interpretation H\<^sub>D\<^sub>'oFF: composite_functor C.VV.comp D'.VV.comp D'.comp FF
\<open>\<lambda>\<mu>\<nu>. D'.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
..
interpretation FoH\<^sub>C: composite_functor C.VV.comp V\<^sub>C D'.comp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>C snd \<mu>\<nu>\<close>
F
..
interpretation natural_transformation C.VV.comp D'.comp H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C.map \<Phi>
proof
show "\<And>\<mu>\<nu>. \<not> C.VV.arr \<mu>\<nu> \<Longrightarrow> \<Phi> \<mu>\<nu> = D'.null"
by (simp add: F.\<Phi>.is_extensional)
fix \<mu>\<nu>
assume \<mu>\<nu>: "C.VV.arr \<mu>\<nu>"
have 1: "D'.arr (\<Phi> \<mu>\<nu>)"
- using \<mu>\<nu> D'.arr_char F.\<Phi>.is_natural_1 F.\<Phi>.components_are_iso
+ using \<mu>\<nu> D'.arr_char\<^sub>S\<^sub>b\<^sub>C F.\<Phi>.is_natural_1 F.\<Phi>.components_are_iso
by (metis (no_types, lifting) D.src_vcomp D.trg_vcomp FoH\<^sub>C.preserves_arr
F.\<Phi>.preserves_reflects_arr)
show "D'.dom (\<Phi> \<mu>\<nu>) = H\<^sub>D\<^sub>'oFF.map (C.VV.dom \<mu>\<nu>)"
- using 1 \<mu>\<nu> D'.dom_char C.VV.arr_char C.VV.dom_char F.FF_def FF_def D'.hcomp_def
+ using 1 \<mu>\<nu> D'.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C F.FF_def FF_def D'.hcomp_def
by simp
show "D'.cod (\<Phi> \<mu>\<nu>) = FoH\<^sub>C.map (C.VV.cod \<mu>\<nu>)"
- using 1 \<mu>\<nu> D'.cod_char C.VV.arr_char F.FF_def FF_def D'.hcomp_def by simp
+ using 1 \<mu>\<nu> D'.cod_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.FF_def FF_def D'.hcomp_def by simp
show "D'.comp (FoH\<^sub>C.map \<mu>\<nu>) (\<Phi> (C.VV.dom \<mu>\<nu>)) = \<Phi> \<mu>\<nu>"
- using 1 \<mu>\<nu> D'.arr_char D'.comp_char C.VV.dom_char F.\<Phi>.is_natural_1
+ using 1 \<mu>\<nu> D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.comp_char C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C F.\<Phi>.is_natural_1
C.VV.arr_dom D.src_vcomp D.trg_vcomp FoH\<^sub>C.preserves_arr F.\<Phi>.preserves_reflects_arr
by (metis (mono_tags, lifting))
show "D'.comp (\<Phi> (C.VV.cod \<mu>\<nu>)) (H\<^sub>D\<^sub>'oFF.map \<mu>\<nu>) = \<Phi> \<mu>\<nu>"
proof -
have "Obj (F.map\<^sub>0 (trg\<^sub>C (fst \<mu>\<nu>))) \<and> Obj (F.map\<^sub>0 (trg\<^sub>C (snd \<mu>\<nu>)))"
- using \<mu>\<nu> C.VV.arr_char
- by (metis (no_types, lifting) C.src_trg C.trg.preserves_reflects_arr D'.arr_char
+ using \<mu>\<nu> C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis (no_types, lifting) C.src_trg C.trg.preserves_reflects_arr D'.arr_char\<^sub>S\<^sub>b\<^sub>C
F.map\<^sub>0_def preserves_hseq)
moreover have "Obj (F.map\<^sub>0 (src\<^sub>C (snd \<mu>\<nu>)))"
- using \<mu>\<nu> C.VV.arr_char
- by (metis (no_types, lifting) C.src.preserves_reflects_arr C.trg_src D'.arr_char
+ using \<mu>\<nu> C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis (no_types, lifting) C.src.preserves_reflects_arr C.trg_src D'.arr_char\<^sub>S\<^sub>b\<^sub>C
F.map\<^sub>0_def preserves_hseq)
ultimately show ?thesis
- using \<mu>\<nu> 1 D'.arr_char D'.comp_char D'.hseq_char C.VV.arr_char C.VV.cod_char
+ using \<mu>\<nu> 1 D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.comp_char D'.hseq_char C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
C.VxV.cod_char FF_def F.FF_def D'.hcomp_char preserves_hseq
apply simp
using F.\<Phi>.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 \<Phi>
apply unfold_locales
- using D'.iso_char D'.arr_char by fastforce
+ using D'.iso_char\<^sub>S\<^sub>b\<^sub>C D'.arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
sublocale pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.hcomp D'.\<a> \<i>\<^sub>D D'.src D'.trg
F \<Phi>
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 \<a>\<^sub>C[f, g, h]) (D'.comp (\<Phi> (f \<star>\<^sub>C g, h)) (D'.hcomp (\<Phi> (f, g)) (F h))) =
F \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)"
proof -
have 1: "D'.arr (cmp (f, g) \<star>\<^sub>D map h)"
- by (metis (mono_tags, lifting) C.ideD(1) D'.arr_char D'.hcomp_closed
+ by (metis (mono_tags, lifting) C.ideD(1) D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.hcomp_closed
F.\<Phi>.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 \<star>\<^sub>C g, h)) (cmp (f, g) \<star>\<^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
+ using 1 f g h fg gh D'.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
by (intro D.seqI) auto
moreover have "D'.arr (cmp (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (cmp (f, g) \<star>\<^sub>D map h))"
- using 1 2 D'.arr_char
- by (metis (no_types, lifting) D'.comp_char D'.seq_char D.seqE F.\<Phi>.preserves_reflects_arr
+ using 1 2 D'.arr_char\<^sub>S\<^sub>b\<^sub>C
+ by (metis (no_types, lifting) D'.comp_char D'.seq_char\<^sub>S\<^sub>b\<^sub>C D.seqE F.\<Phi>.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
+ using f g h fg gh D'.dom_char\<^sub>S\<^sub>b\<^sub>C D'.cod_char\<^sub>S\<^sub>b\<^sub>C D'.comp_char D'.hcomp_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply simp
by force
qed
also have "... = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^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 (\<Phi> (f, g \<star>\<^sub>C h))
(D'.comp (D'.hcomp (F f) (\<Phi> (g, h))) (D'.\<a> (F f) (F g) (F h)))"
proof -
have "D.seq (map f \<star>\<^sub>D cmp (g, h)) \<a>\<^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
+ using f g h fg gh C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
by (intro D.seqI) auto
moreover have "D'.arr \<a>\<^sub>D[map f, map g, map h]"
- using f g h fg gh D'.arr_char preserves_arr by auto
+ using f g h fg gh D'.arr_char\<^sub>S\<^sub>b\<^sub>C preserves_arr by auto
moreover have "D'.arr (map f \<star>\<^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)
+ by (metis (no_types, lifting) D'.arr_char\<^sub>S\<^sub>b\<^sub>C D.seqE D.vseq_implies_hpar(1)
D.vseq_implies_hpar(2) calculation(1-2))
moreover have "D'.arr ((map f \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D \<a>\<^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
+ by (metis (no_types, lifting) D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.comp_closed D.seqE
calculation(1-3))
moreover have "D.seq (cmp (f, g \<star>\<^sub>C h))
((map f \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D \<a>\<^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
+ using f g h fg gh C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D'.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C D'.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D'.comp_char
D'.hcomp_def
by simp
qed
finally show "D'.comp (F \<a>\<^sub>C[f, g, h])
(D'.comp (\<Phi> (f \<star>\<^sub>C g, h)) (D'.hcomp (\<Phi> (f, g)) (F h))) =
D'.comp (\<Phi> (f, g \<star>\<^sub>C h))
(D'.comp (D'.hcomp (F f) (\<Phi> (g, h))) (D'.\<a> (F f) (F g) (F h)))"
by blast
qed
lemma is_pseudofunctor:
shows "pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.hcomp D'.\<a> \<i>\<^sub>D D'.src D'.trg F \<Phi>"
..
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)"
- using D'.arr_char D'.in_hom_char assms unit_in_hom(2) by force
+ using D'.arr_char\<^sub>S\<^sub>b\<^sub>C D'.in_hom_char\<^sub>S\<^sub>b\<^sub>C 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)
+ using assms D'.iso_char\<^sub>S\<^sub>b\<^sub>C D'.arr_char\<^sub>S\<^sub>b\<^sub>C F.unit_char(2)
\<open>D'.in_hom (F.unit a) (map\<^sub>0 a) (map a)\<close>
by auto
show "D'.comp (F.unit a) \<i>\<^sub>D[map\<^sub>0 a] =
D'.comp (D'.comp (map \<i>\<^sub>C[a]) (cmp (a, a)))
(D'.hcomp (F.unit a) (F.unit a))"
proof -
have "D'.comp (F.unit a) \<i>\<^sub>D[map\<^sub>0 a] = F.unit a \<cdot>\<^sub>D \<i>\<^sub>D[src\<^sub>D (map a)]"
- using assms D'.comp_char D'.arr_char
+ using assms D'.comp_char D'.arr_char\<^sub>S\<^sub>b\<^sub>C
apply simp
by (metis (no_types, lifting) C.obj_simps(1-2) F.preserves_src preserves_arr)
also have "... = (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.unit a)"
using assms F.unit_char(3) [of a] by auto
also have "... = D'.comp (D'.comp (map \<i>\<^sub>C[a]) (cmp (a, a)))
(D'.hcomp (F.unit a) (F.unit a))"
proof -
have "D'.arr (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a))"
using assms D'.comp_simp by auto
moreover have "D.seq (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a)) (F.unit a \<star>\<^sub>D F.unit a)"
- using assms C.VV.arr_char F.cmp_simps(4-5)
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C F.cmp_simps(4-5)
by (intro D.seqI) auto
ultimately show ?thesis
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)
+ D'.iso_is_arr D'.seq_char\<^sub>S\<^sub>b\<^sub>C D'.vconn_implies_hpar(1-2)
\<i>_simps(1) \<open>D'.iso (F.unit a)\<close> assms map\<^sub>0_simps(2-3))
qed
finally show ?thesis by blast
qed
qed
end
subsection "Equivalence Pseudofunctors"
text \<open>
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.
\<close>
context pseudofunctor
begin
definition locally_faithful
where "locally_faithful \<equiv>
\<forall>f g \<mu> \<mu>'. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C g\<guillemotright> \<and> \<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>C g\<guillemotright> \<and> F \<mu> = F \<mu>' \<longrightarrow> \<mu> = \<mu>'"
lemma locally_faithful_iff_faithful:
shows "locally_faithful \<longleftrightarrow> faithful_functor V\<^sub>C V\<^sub>D F"
proof
show "faithful_functor V\<^sub>C V\<^sub>D F \<Longrightarrow> locally_faithful"
by (metis category.in_homE faithful_functor.is_faithful functor_axioms
functor_def locally_faithful_def)
show "locally_faithful \<Longrightarrow> 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 \<mu> \<mu>'
assume par: "C.par \<mu> \<mu>'" and eq: "F \<mu> = F \<mu>'"
obtain f g where fg: "\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C g\<guillemotright> \<and> \<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>C g\<guillemotright>"
using par by auto
show "\<mu> = \<mu>'"
using 1 fg locally_faithful_def eq by simp
qed
qed
qed
end
text \<open>
In contrast, it is not true that a pseudofunctor that is locally full is also
globally full, because we can have \<open>\<guillemotleft>\<nu> : F h \<Rightarrow>\<^sub>D F k\<guillemotright>\<close> even if \<open>h\<close> and \<open>k\<close>
are not in the same hom-category. So it would be a mistake to require that an
equivalence functor be globally full.
\<close>
locale equivalence_pseudofunctor =
pseudofunctor +
faithful_functor V\<^sub>C V\<^sub>D F +
assumes biessentially_surjective_on_objects:
"D.obj a' \<Longrightarrow> \<exists>a. C.obj a \<and> D.equivalent_objects (map\<^sub>0 a) a'"
and locally_essentially_surjective:
"\<lbrakk> C.obj a; C.obj b; \<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 b\<guillemotright>; D.ide g \<rbrakk> \<Longrightarrow>
\<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) g"
and locally_full:
"\<lbrakk> C.ide f; C.ide f'; src\<^sub>C f = src\<^sub>C f'; trg\<^sub>C f = trg\<^sub>C f'; \<guillemotleft>\<nu> : F f \<Rightarrow>\<^sub>D F f'\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> F \<mu> = \<nu>"
begin
lemma reflects_ide:
assumes "C.endo \<mu>" and "D.ide (F \<mu>)"
shows "C.ide \<mu>"
using assms is_faithful [of "C.dom \<mu>" \<mu>] 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 \<mu>" and "D.iso (F \<mu>)"
shows "C.iso \<mu>"
proof -
obtain \<mu>' where \<mu>': "\<guillemotleft>\<mu>' : C.cod \<mu> \<Rightarrow>\<^sub>C C.dom \<mu>\<guillemotright> \<and> F \<mu>' = D.inv (F \<mu>)"
using assms locally_full [of "C.cod \<mu>" "C.dom \<mu>" "D.inv (F \<mu>)"]
D.inv_in_hom C.in_homE preserves_hom C.in_homI
by auto
have "C.inverse_arrows \<mu> \<mu>'"
using assms \<mu>' 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 "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>" and "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>"
and "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
shows "equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
proof -
interpret E': equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close>
\<open>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)\<close>
\<open>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)\<close>
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 "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>"
using assms(3) by simp
show "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>"
using assms(4) by simp
have 0: "src\<^sub>C f = trg\<^sub>C g \<and> src\<^sub>C g = trg\<^sub>C f"
using \<open>\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>\<close>
by (metis C.hseqE C.ideD(1) C.ide_cod C.ide_dom C.in_homE assms(4))
show "C.iso \<eta>"
proof -
have "D.iso (F \<eta>)"
proof -
have 1: "\<guillemotleft>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) : map\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F f\<guillemotright>"
using \<open>C.ide f\<close> \<open>C.ide g\<close> \<open>\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>\<close>
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 (\<Phi> (g, f)) \<and> \<guillemotleft>\<Phi> (g, f) : F g \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C f)\<guillemotright>"
using 0 \<open>C.ide f\<close> \<open>C.ide g\<close> cmp_in_hom by simp
have 3: "D.iso (D.inv (unit (src\<^sub>C f))) \<and>
\<guillemotleft>D.inv (unit (src\<^sub>C f)) : F (src\<^sub>C f) \<Rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C f)\<guillemotright>"
using \<open>C.ide f\<close> unit_char by simp
have "D.iso (\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f)))"
using 1 2 3 E'.unit_is_iso D.isos_compose by blast
moreover have "\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f)) =
F \<eta>"
proof -
have "\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f))
= (\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D (unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f)))"
using D.comp_assoc by simp
also have "... = F (g \<star>\<^sub>C f) \<cdot>\<^sub>D F \<eta> \<cdot>\<^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 \<eta>"
using \<open>\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>\<close> 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 \<epsilon>"
proof -
have "D.iso (F \<epsilon>)"
proof -
have 1: "\<guillemotleft>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) : F f \<star>\<^sub>D F g \<Rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C g)\<guillemotright>"
using \<open>C.ide f\<close> \<open>C.ide g\<close> \<open>\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>\<close>
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 (\<Phi> (f, g))) \<and>
\<guillemotleft>D.inv (\<Phi> (f, g)) : F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F g\<guillemotright>"
using 0 \<open>C.ide f\<close> \<open>C.ide g\<close> \<open>\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>\<close> cmp_in_hom(2) D.inv_in_hom
by simp
have 3: "D.iso (unit (trg\<^sub>C f)) \<and> \<guillemotleft>unit (trg\<^sub>C f) : map\<^sub>0 (trg\<^sub>C f) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
using \<open>C.ide f\<close> unit_char by simp
have "D.iso (unit (trg\<^sub>C f) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (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) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g)) =
F \<epsilon>"
proof -
have "unit (trg\<^sub>C f) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g)) =
(unit (trg\<^sub>C f) \<cdot>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D (\<Phi> (f, g) \<cdot>\<^sub>D D.inv (\<Phi> (f, g)))"
using D.comp_assoc by simp
also have "... = F (trg\<^sub>C f) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D F (f \<star>\<^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))
+ by (simp add: C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C assms(1-2))
also have "... = F \<epsilon>"
using 0 \<open>\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>\<close> 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' \<phi> \<psi> where E': "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) g' \<phi> \<psi>"
using assms D.equivalence_map_def by auto
interpret E': equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> g' \<phi> \<psi>
using E' by auto
obtain g where g: "\<guillemotleft>g : trg\<^sub>C f \<rightarrow>\<^sub>C src\<^sub>C f\<guillemotright> \<and> C.ide g \<and> 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 \<mu> where \<mu>: "\<guillemotleft>\<mu> : g' \<Rightarrow>\<^sub>D F g\<guillemotright> \<and> D.iso \<mu>"
using g D.isomorphic_def D.isomorphic_symmetric by blast
interpret E'': equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close>
\<open>(F g \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D F f) \<cdot>\<^sub>D \<phi>\<close>
\<open>\<psi> \<cdot>\<^sub>D (D.inv (F f) \<star>\<^sub>D g') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<mu>)\<close>
using assms \<mu> E'.equivalence_in_bicategory_axioms D.ide_is_iso
D.equivalence_respects_iso [of "F f" g' \<phi> \<psi> "F f" "F f" \<mu> "F g"]
by auto
let ?\<eta>' = "\<Phi> (g, f) \<cdot>\<^sub>D (F g \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D F f) \<cdot>\<^sub>D \<phi> \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
have \<eta>': "\<guillemotleft>?\<eta>' : F (src\<^sub>C f) \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C f)\<guillemotright>"
using assms \<mu> 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_\<eta>': "D.iso ?\<eta>'"
using assms g \<mu> \<eta>' 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 ?\<epsilon>' = "unit (src\<^sub>C g) \<cdot>\<^sub>D \<psi> \<cdot>\<^sub>D (D.inv (F f) \<star>\<^sub>D g') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<mu>) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g))"
have \<epsilon>': "\<guillemotleft>?\<epsilon>' : F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>D.inv (\<Phi> (f, g)) : F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F g\<guillemotright>"
- using assms g cmp_in_hom C.VV.ide_char C.VV.arr_char by auto
+ using assms g cmp_in_hom C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
show "\<guillemotleft>F f \<star>\<^sub>D D.inv \<mu> : F f \<star>\<^sub>D F g \<Rightarrow>\<^sub>D F f \<star>\<^sub>D g'\<guillemotright>"
using assms g \<mu> E''.antipar D.ide_in_hom(2) by auto
show "\<guillemotleft>D.inv (F f) \<star>\<^sub>D g' : F f \<star>\<^sub>D g' \<Rightarrow>\<^sub>D F f \<star>\<^sub>D g'\<guillemotright>"
using assms E'.antipar D.ide_is_iso by auto
show "\<guillemotleft>\<psi> : F f \<star>\<^sub>D g' \<Rightarrow>\<^sub>D trg\<^sub>D (F f)\<guillemotright>"
using E'.counit_in_hom by simp
show "\<guillemotleft>unit (src\<^sub>C g) : trg\<^sub>D (F f) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
using assms g unit_char by auto
qed
have iso_\<epsilon>': "D.iso ?\<epsilon>'"
proof -
have "D.iso (\<Phi> (f, g))"
- using assms g C.VV.ide_char C.VV.arr_char by auto
+ using assms g C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis
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
\<epsilon>' \<mu> g unit_char(2))
qed
obtain \<eta> where \<eta>: "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright> \<and> F \<eta> = ?\<eta>'"
using assms g E'.antipar \<eta>' locally_full [of "src\<^sub>C f" "g \<star>\<^sub>C f" ?\<eta>']
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_\<eta>: "C.iso \<eta>"
using \<eta> \<eta>' iso_\<eta>' reflects_iso by auto
have 1: "\<exists>\<epsilon>. \<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright> \<and> F \<epsilon> = ?\<epsilon>'"
using assms g \<epsilon>' locally_full [of "f \<star>\<^sub>C g" "src\<^sub>C g" ?\<epsilon>'] by force
obtain \<epsilon> where \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright> \<and> F \<epsilon> = ?\<epsilon>'"
using 1 by blast
have iso_\<epsilon>: "C.iso \<epsilon>"
using \<epsilon> \<epsilon>' iso_\<epsilon>' reflects_iso by auto
have "equivalence_in_bicategory (\<cdot>\<^sub>C) (\<star>\<^sub>C) \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
using assms g \<eta> \<epsilon> iso_\<eta> iso_\<epsilon> 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': "\<guillemotleft>f' : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 b\<guillemotright> \<and> D.equivalence_map f'"
using assms D.equivalent_objects_def D.equivalence_map_def by auto
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright> \<and> C.ide f \<and> 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\<open>
For each pair of objects \<open>a\<close>, \<open>b\<close> of \<open>C\<close>, an equivalence pseudofunctor restricts
to an equivalence of categories between \<open>C.hhom a b\<close> and \<open>D.hhom (map\<^sub>0 a) (map\<^sub>0 b)\<close>.
\<close>
(* 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 \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : a \<rightarrow>\<^sub>C a'\<guillemotright>\<close>
using C.hhom_is_subcategory by simp
sublocale hhom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
definition F\<^sub>1
where "F\<^sub>1 = (\<lambda>\<mu>. if hhom\<^sub>C.arr \<mu> then F \<mu> 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
+ using hhom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>D.arr_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>D.dom_char\<^sub>S\<^sub>b\<^sub>C
+ hhom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>D.cod_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>C.seq_char\<^sub>S\<^sub>b\<^sub>C 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 "\<And>\<mu> \<mu>'. \<lbrakk>hhom\<^sub>C.par \<mu> \<mu>'; F\<^sub>1 \<mu> = F\<^sub>1 \<mu>'\<rbrakk> \<Longrightarrow> \<mu> = \<mu>'"
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
+ using is_faithful hhom\<^sub>C.dom_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>D.dom_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>C.cod_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>D.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis C.in_hhom_def hhom\<^sub>C.arrE)
show "\<And>f f' \<nu>. \<lbrakk>hhom\<^sub>C.ide f; hhom\<^sub>C.ide f'; hhom\<^sub>D.in_hom \<nu> (F\<^sub>1 f') (F\<^sub>1 f)\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. hhom\<^sub>C.in_hom \<mu> f' f \<and> F\<^sub>1 \<mu> = \<nu>"
proof (unfold F\<^sub>1_def)
fix f f' \<nu>
assume f: "hhom\<^sub>C.ide f" and f': "hhom\<^sub>C.ide f'"
assume "hhom\<^sub>D.in_hom \<nu> (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 \<nu>: "hhom\<^sub>D.in_hom \<nu> (F f') (F f)"
using f f' by simp
have "\<exists>\<mu>. hhom\<^sub>C.in_hom \<mu> f' f \<and> F \<mu> = \<nu>"
proof -
have 1: "src\<^sub>C f' = src\<^sub>C f \<and> 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: "\<exists>\<mu>. C.in_hom \<mu> f' f \<and> F \<mu> = \<nu>"
- by (meson \<nu> f f' hhom\<^sub>D.in_hom_char horizontal_homs.hhom_is_subcategory
- locally_full subcategory.ide_char weak_arrow_of_homs_axioms
+ by (meson \<nu> f f' hhom\<^sub>D.in_hom_char\<^sub>S\<^sub>b\<^sub>C horizontal_homs.hhom_is_subcategory
+ locally_full subcategory.ide_char\<^sub>S\<^sub>b\<^sub>C weak_arrow_of_homs_axioms
weak_arrow_of_homs_def)
obtain \<mu> where \<mu>: "C.in_hom \<mu> f' f \<and> F \<mu> = \<nu>"
using ex by blast
have "hhom\<^sub>C.in_hom \<mu> f' f"
by (metis C.arrI C.in_hhom_def C.vconn_implies_hpar(1-2) \<mu> f f'
- hhom\<^sub>C.arr_char hhom\<^sub>C.ide_char hhom\<^sub>C.in_hom_char)
+ hhom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>C.ide_char hhom\<^sub>C.in_hom_char\<^sub>S\<^sub>b\<^sub>C)
thus ?thesis
using \<mu> by auto
qed
thus "\<exists>\<mu>. hhom\<^sub>C.in_hom \<mu> f' f \<and> (if hhom\<^sub>C.arr \<mu> then F \<mu> else D.null) = \<nu>"
by auto
qed
show "\<And>g. hhom\<^sub>D.ide g \<Longrightarrow> \<exists>f. hhom\<^sub>C.ide f \<and> 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 "\<exists>f. hhom\<^sub>C.ide f \<and> hhom\<^sub>D.isomorphic (if hhom\<^sub>C.arr f then F f else D.null) g"
proof -
have "C.obj a \<and> C.obj a'"
using obj_a obj_a' by simp
moreover have 1: "D.ide g \<and> \<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>"
- using g obj_a obj_a' hhom\<^sub>D.ide_char by auto
+ using g obj_a obj_a' hhom\<^sub>D.ide_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately have 2: "\<exists>f. C.in_hhom f a a' \<and> C.ide f \<and> D.isomorphic (F f) g"
using locally_essentially_surjective [of a a' g] by simp
obtain f \<phi> where f: "C.in_hhom f a a' \<and> C.ide f \<and> D.in_hom \<phi> (F f) g \<and> D.iso \<phi>"
using 2 by auto
have "hhom\<^sub>C.ide f"
- using f hhom\<^sub>C.ide_char hhom\<^sub>C.arr_char by simp
+ using f hhom\<^sub>C.ide_char\<^sub>S\<^sub>b\<^sub>C hhom\<^sub>C.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
moreover have "hhom\<^sub>D.isomorphic (F f) g"
proof -
have "hhom\<^sub>D.arr \<phi> \<and> hhom\<^sub>D.arr (D.inv \<phi>)"
by (metis 1 D.arrI D.in_hhom_def D.vconn_implies_hpar(1-4) D.inv_in_homI
- f hhom\<^sub>D.arrI)
+ f hhom\<^sub>D.arrI\<^sub>S\<^sub>b\<^sub>C)
hence "hhom\<^sub>D.in_hom \<phi> (F f) g \<and> hhom\<^sub>D.iso \<phi>"
- 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)
+ 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\<^sub>S\<^sub>b\<^sub>C)
thus ?thesis
unfolding hhom\<^sub>D.isomorphic_def by blast
qed
ultimately show "\<exists>f. hhom\<^sub>C.ide f \<and>
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. \<exists>\<eta>\<epsilon>.
adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>))"
lemma G\<^sub>1_props:
assumes "C.obj a" and "C.obj a'"
shows "\<exists>\<eta> \<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"
proof -
have "\<exists>G. \<exists>\<eta>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>)"
using F\<^sub>1.extends_to_adjoint_equivalence by simp
hence "\<exists>\<eta>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>)"
unfolding G\<^sub>1_def
using someI_ex
[of "\<lambda>G. \<exists>\<eta>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>)"]
by blast
thus ?thesis by simp
qed
definition \<eta>
where "\<eta> = (SOME \<eta>. \<exists>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>)"
definition \<epsilon>
where "\<epsilon> = (SOME \<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>)"
lemma \<eta>\<epsilon>_props:
shows "adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"
using obj_a obj_a' \<eta>_def \<epsilon>_def G\<^sub>1_props
someI_ex [of "\<lambda>\<eta>. \<exists>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"]
by simp
sublocale \<eta>\<epsilon>: adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>
using \<eta>\<epsilon>_props by simp
sublocale \<eta>\<epsilon>: meta_adjunction hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta>\<epsilon>.\<phi> \<eta>\<epsilon>.\<psi>
using \<eta>\<epsilon>.induces_meta_adjunction by simp
end
context identity_pseudofunctor
begin
sublocale equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
map cmp
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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F +
G: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
begin
interpretation faithful_functor V\<^sub>B V\<^sub>D \<open>G o F\<close>
using F.faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
by blast
interpretation equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>G o F\<close> cmp
proof
show "\<And>c. D.obj c \<Longrightarrow> \<exists>a. B.obj a \<and> D.equivalent_objects (map\<^sub>0 a) c"
proof -
fix c
assume c: "D.obj c"
obtain b where b: "C.obj b \<and> 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 \<and> 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 "\<exists>a. B.obj a \<and> D.equivalent_objects (map\<^sub>0 a) c"
using a by auto
qed
show "\<And>a a' h. \<lbrakk>B.obj a; B.obj a'; \<guillemotleft>h : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>; D.ide h\<rbrakk>
\<Longrightarrow> \<exists>f. B.in_hhom f a a' \<and> B.ide f \<and> D.isomorphic ((G \<circ> F) f) h"
proof -
fix a a' h
assume a: "B.obj a" and a': "B.obj a'"
and h_in_hom: "\<guillemotleft>h : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>" and ide_h: "D.ide h"
obtain g
where g: "C.in_hhom g (F.map\<^sub>0 a) (F.map\<^sub>0 a') \<and> C.ide g \<and> 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' \<and> B.ide f \<and> C.isomorphic (F f) g"
using a a' g F.locally_essentially_surjective by blast
have "D.isomorphic ((G o F) f) h"
by (metis D.isomorphic_transitive G.preserves_isomorphic comp_apply f g)
thus "\<exists>f. B.in_hhom f a a' \<and> B.ide f \<and> D.isomorphic ((G \<circ> F) f) h"
using f by auto
qed
show "\<And>f f' \<nu>. \<lbrakk>B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f';
\<guillemotleft>\<nu> : (G \<circ> F) f \<Rightarrow>\<^sub>D (G \<circ> F) f'\<guillemotright>\<rbrakk>
\<Longrightarrow> \<exists>\<tau>. \<guillemotleft>\<tau> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> (G \<circ> F) \<tau> = \<nu>"
proof -
fix f f' \<nu>
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 \<nu>: "\<guillemotleft>\<nu> : (G \<circ> F) f \<Rightarrow>\<^sub>D (G \<circ> F) f'\<guillemotright>"
have \<nu>: "\<guillemotleft>\<nu> : G (F f) \<Rightarrow>\<^sub>D G (F f')\<guillemotright>"
using \<nu> by simp
have 1: "src\<^sub>C (F f) = src\<^sub>C (F f') \<and> trg\<^sub>C (F f) = trg\<^sub>C (F f')"
using f f' src trg by simp
have 2: "\<exists>\<mu>. \<guillemotleft>\<mu> : F f \<Rightarrow>\<^sub>C F f'\<guillemotright> \<and> G \<mu> = \<nu>"
using f f' 1 \<nu> G.locally_full F.preserves_ide by simp
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : F f \<Rightarrow>\<^sub>C F f'\<guillemotright> \<and> G \<mu> = \<nu>"
using 2 by auto
obtain \<tau> where \<tau>: "\<guillemotleft>\<tau> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> F \<tau> = \<mu>"
using f f' src trg 2 \<mu> F.locally_full by blast
show "\<exists>\<tau>. \<guillemotleft>\<tau> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> (G \<circ> F) \<tau> = \<nu>"
using \<mu> \<tau> by auto
qed
qed
sublocale equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>G o F\<close> cmp ..
lemma is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(G o F) cmp"
..
end
end
diff --git a/thys/Bicategory/PseudonaturalTransformation.thy b/thys/Bicategory/PseudonaturalTransformation.thy
--- a/thys/Bicategory/PseudonaturalTransformation.thy
+++ b/thys/Bicategory/PseudonaturalTransformation.thy
@@ -1,8202 +1,8202 @@
(* Title: PseudonaturalTransformation
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2020
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Pseudonatural Transformations"
theory PseudonaturalTransformation
imports InternalEquivalence InternalAdjunction Pseudofunctor
begin
subsection "Definition of Pseudonatural Transformation"
text \<open>
Pseudonatural transformations are a generalization of natural transformations that is
appropriate for pseudofunctors. The ``components'' of a pseudonatural transformation \<open>\<tau>\<close>
from a pseudofunctor \<open>F\<close> to a pseudofunctor \<open>G\<close> (both from bicategory \<open>C\<close> to \<open>D\<close>),
are 1-cells \<open>\<guillemotleft>\<tau>\<^sub>0 a : F\<^sub>0 a \<rightarrow>\<^sub>D G\<^sub>0 a\<guillemotright>\<close> associated with the objects of \<open>C\<close>. Instead of
``naturality squares'' that commute on-the-nose, a pseudonatural transformation associates
an invertible 2-cell \<open>\<guillemotleft>\<tau>\<^sub>1 f : G f \<star>\<^sub>D \<tau>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0 a' \<star>\<^sub>D F f\<guillemotright>\<close> with each 1-cell \<open>\<guillemotleft>f : a \<rightarrow>\<^sub>C a'\<guillemotright>\<close>
of \<open>C\<close>. The 1-cells \<open>\<tau>\<^sub>0 a\<close> and 2-cells \<open>\<tau>\<^sub>1 f\<close> are subject to coherence conditions which
express that they transform naturally across 2-cells of \<open>C\<close> and behave sensibly with
respect to objects and horizontal composition.
In formalizing ordinary natural transformations, we found it convenient to treat them
similarly to functors in that a natural transformation \<open>\<tau> : C \<rightarrow> D\<close> maps arrows of \<open>C\<close>
to arrows of \<open>D\<close>; the components \<open>\<tau> a\<close> at objects \<open>a\<close> being merely special cases.
However, it is not possible to take the same approach for pseudofunctors, because it is
not possible to treat the components \<open>\<tau>\<^sub>0 a\<close> at objects \<open>a\<close> as a special case of the
components \<open>\<tau>\<^sub>1 f\<close> at 1-cells \<open>f\<close>. So we have to regard a pseudonatural transformation \<open>\<tau>\<close>
as consisting of two separate mappings: a mapping \<open>\<tau>\<^sub>0\<close> from objects of \<open>C\<close> to 1-cells
of \<open>D\<close> and a mapping \<open>\<tau>\<^sub>1\<close> from 1-cells of \<open>C\<close> to invertible 2-cells of \<open>D\<close>.
Pseudonatural transformations are themselves a special case of the more general notion
of ``lax natural transformations'' between pseudofunctors. For a lax natural transformation
\<open>\<tau>\<close>, the 2-cells \<open>\<tau>\<^sub>1 f\<close> are not required to be invertible. This means that there is a
distinction between a lax natural transformation with \<open>\<guillemotleft>\<tau>\<^sub>1 f : G f \<star>\<^sub>D \<tau>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0 a' \<star>\<^sub>D F f\<guillemotright>\<close>
and an ``op-lax'' natural transformation with \<open>\<guillemotleft>\<tau>\<^sub>1 f : \<tau>\<^sub>0 a' \<star>\<^sub>D F f \<Rightarrow>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 a\<guillemotright>\<close>.
There is some variation in the literature on which direction is considered ``lax'' and
which is ``op-lax'' and this variation extends as well to the special case of pseudofunctors,
though in that case it does not result in any essential difference in the notion,
due to the assumed invertibility. We have chosen the direction that agrees with
Leinster \cite{leinster-basic-bicategories}, and with the ``nLab'' article
\cite{nlab-lax-natural-transformation} on lax natural transformations, but note that the
``nLab'' article \cite{nlab-pseudonatural-transformation} on pseudonatural transformations
seems to make the opposite choice.
\<close>
locale pseudonatural_transformation =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F +
G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and \<tau>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1 :: "'c \<Rightarrow> 'd" +
assumes map\<^sub>0_in_hhom: "C.obj a \<Longrightarrow> \<guillemotleft>\<tau>\<^sub>0 a : src\<^sub>D (F a) \<rightarrow>\<^sub>D src\<^sub>D (G a)\<guillemotright>"
and map\<^sub>1_in_vhom: "C.ide f \<Longrightarrow> \<guillemotleft>\<tau>\<^sub>1 f : G f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F f\<guillemotright>"
and ide_map\<^sub>0_obj: "C.obj a \<Longrightarrow> D.ide (\<tau>\<^sub>0 a)"
and iso_map\<^sub>1_ide: "C.ide f \<Longrightarrow> D.iso (\<tau>\<^sub>1 f)"
and naturality: "C.arr \<mu> \<Longrightarrow>
\<tau>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) = (\<tau>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 (C.dom \<mu>)"
and respects_unit: "C.obj a \<Longrightarrow>
(\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a] = \<tau>\<^sub>1 a \<cdot>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
and respects_hcomp:
"\<lbrakk> C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f \<rbrakk> \<Longrightarrow>
(\<tau>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D (\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G g, \<tau>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 (src\<^sub>C f)]
= \<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))"
begin
lemma map\<^sub>0_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>\<tau>\<^sub>0 a : F.map\<^sub>0 a \<rightarrow>\<^sub>D G.map\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<tau>\<^sub>0 a : \<tau>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0 a\<guillemotright>"
using assms map\<^sub>0_in_hhom [of a]
apply fastforce
using assms ide_map\<^sub>0_obj [of a] by fast
lemma map\<^sub>0_simps [simp]:
assumes "C.obj a"
shows "D.ide (\<tau>\<^sub>0 a)"
and "src\<^sub>D (\<tau>\<^sub>0 a) = F.map\<^sub>0 a" and "trg\<^sub>D (\<tau>\<^sub>0 a) = G.map\<^sub>0 a"
using assms map\<^sub>0_in_hom iso_map\<^sub>1_ide ide_map\<^sub>0_obj by auto
lemma map\<^sub>1_in_hom [intro]:
assumes "C.ide f"
shows "\<guillemotleft>\<tau>\<^sub>1 f : F.map\<^sub>0 (src\<^sub>C f) \<rightarrow>\<^sub>D G.map\<^sub>0 (trg\<^sub>C f)\<guillemotright>"
and "\<guillemotleft>\<tau>\<^sub>1 f : G f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F f\<guillemotright>"
using assms map\<^sub>1_in_vhom [of f] D.vconn_implies_hpar(1-2) by auto
lemma map\<^sub>1_simps [simp]:
assumes "C.ide f"
shows "D.arr (\<tau>\<^sub>1 f)"
and "src\<^sub>D (\<tau>\<^sub>1 f) = F.map\<^sub>0 (src\<^sub>C f)"
and "trg\<^sub>D (\<tau>\<^sub>1 f) = G.map\<^sub>0 (trg\<^sub>C f)"
and "D.dom (\<tau>\<^sub>1 f) = G f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)"
and "D.cod (\<tau>\<^sub>1 f) = \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F f"
using assms map\<^sub>1_in_hom iso_map\<^sub>1_ide by auto
lemma inv_naturality:
assumes "C.arr \<mu>"
shows "(G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 (C.dom \<mu>)) =
D.inv (\<tau>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D (\<tau>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>)"
using assms naturality iso_map\<^sub>1_ide D.invert_opposite_sides_of_square by simp
end
subsection "Identity Pseudonatural Transformation"
locale identity_pseudonatural_transformation =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
begin
abbreviation (input) map\<^sub>0
where "map\<^sub>0 a \<equiv> F.map\<^sub>0 a"
abbreviation (input) map\<^sub>1
where "map\<^sub>1 f \<equiv> \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
sublocale pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F F \<Phi>\<^sub>F map\<^sub>0 map\<^sub>1
proof
show 1: "\<And>a. C.obj a \<Longrightarrow> D.ide (F.map\<^sub>0 a)"
using F.map\<^sub>0_simps(1) by blast
show "\<And>f. C.ide f \<Longrightarrow> D.iso (map\<^sub>1 f)"
by auto
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>F.map\<^sub>0 a : src\<^sub>D (F a) \<rightarrow>\<^sub>D src\<^sub>D (F a)\<guillemotright>"
by (metis C.obj_def' D.in_hhomI D.src.preserves_arr F.map\<^sub>0_def F.map\<^sub>0_simps(2)
F.map\<^sub>0_simps(3) F.preserves_reflects_arr)
show "\<And>f. C.ide f \<Longrightarrow> \<guillemotleft>map\<^sub>1 f : F f \<star>\<^sub>D F.map\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D F.map\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F f\<guillemotright>"
by simp
show "\<And>\<mu>. C.arr \<mu> \<Longrightarrow>
(map\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D F.map\<^sub>0 (src\<^sub>C \<mu>))
= (F.map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D map\<^sub>1 (C.dom \<mu>)"
by (metis D.comp_assoc D.lunit'_naturality D.runit_naturality F.preserves_arr
F.preserves_cod F.preserves_dom F.preserves_src F.preserves_trg)
show "\<And>f g. \<lbrakk>C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\<rbrakk> \<Longrightarrow>
(F.map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<cdot>\<^sub>D \<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)]
= (\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D \<r>\<^sub>D[F (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D F.map\<^sub>0 (src\<^sub>C f))"
proof -
fix f g
assume f: "C.ide f" and g: "C.ide g" and fg: "src\<^sub>C g = trg\<^sub>C f"
have "(F.map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<cdot>\<^sub>D \<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)]
= (F.map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
(\<a>\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(\<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
(D.inv \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f])) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<r>\<^sub>D[F f]) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)]"
using f g fg D.whisker_right [of "F f"] D.whisker_left [of "F g"] D.comp_assoc
by simp
also have "... = (F.map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<cdot>\<^sub>D
((\<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f)) \<cdot>\<^sub>D
\<r>\<^sub>D[F g \<star>\<^sub>D F f]"
proof -
have "D.inv \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]) = \<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f"
proof -
have "D.inv \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]) =
D.inv ((F g \<star>\<^sub>D \<l>\<^sub>D[F f]) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f])"
using f g fg 1 D.inv_comp by auto
also have "... = D.inv (\<r>\<^sub>D[F g] \<star>\<^sub>D F f)"
using f g fg D.triangle [of "F f" "F g"] by simp
also have "... = \<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f"
using f g fg by simp
finally show ?thesis by blast
qed
moreover have "(F g \<star>\<^sub>D \<r>\<^sub>D[F f]) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)] = \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
using f g fg D.runit_hcomp(1) by simp
moreover have "\<a>\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f]"
using f g fg D.lunit_hcomp(4) by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((F.map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f]) \<cdot>\<^sub>D \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
proof -
have "((\<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<r>\<^sub>D[F g \<star>\<^sub>D F f] = \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
proof -
have "((\<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<r>\<^sub>D[F g \<star>\<^sub>D F f] =
(\<r>\<^sub>D[F g] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
using f g fg D.whisker_right by simp
also have "... = (F g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
using f g fg D.comp_arr_inv' by simp
also have "... = \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
using f g fg D.comp_cod_arr by simp
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D \<Phi>\<^sub>F (g, f) \<cdot>\<^sub>D \<r>\<^sub>D[F g \<star>\<^sub>D F f]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.lunit'_naturality [of "\<Phi>\<^sub>F (g, f)"] D.comp_assoc
by simp
also have "... = (\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D \<r>\<^sub>D[F (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D F.map\<^sub>0 (src\<^sub>C f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.runit_naturality [of "\<Phi>\<^sub>F (g, f)"] D.comp_assoc
by simp
finally show "(F.map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<cdot>\<^sub>D \<r>\<^sub>D[F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<cdot>\<^sub>D \<r>\<^sub>D[F f]) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)]
= (\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D \<r>\<^sub>D[F (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D F.map\<^sub>0 (src\<^sub>C f))"
by simp
qed
show "\<And>a. C.obj a \<Longrightarrow>
(F.map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[F.map\<^sub>0 a] =
map\<^sub>1 a \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.map\<^sub>0 a)"
proof -
fix a
assume a: "C.obj a"
have "(F.map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[F.map\<^sub>0 a] =
(F.map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[F.map\<^sub>0 a]"
using a 1 D.unitor_coincidence by simp
also have "... = ((F.map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a]) \<cdot>\<^sub>D \<r>\<^sub>D[F.map\<^sub>0 a]"
using D.comp_assoc by simp
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[F a] \<cdot>\<^sub>D F.unit a \<cdot>\<^sub>D \<r>\<^sub>D[F.map\<^sub>0 a]"
using a 1 D.lunit'_naturality [of "F.unit a"] D.comp_assoc by simp
also have "... = map\<^sub>1 a \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.map\<^sub>0 a)"
using a 1 D.runit_naturality [of "F.unit a"] D.comp_assoc by simp
finally show "(F.map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[F.map\<^sub>0 a] =
map\<^sub>1 a \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.map\<^sub>0 a)"
by blast
qed
qed
lemma is_pseudonatural_transformation:
shows "pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F F \<Phi>\<^sub>F map\<^sub>0 map\<^sub>1"
..
end
subsection "Composite Pseudonatural Transformation"
text \<open>
A pseudonatural transformation \<open>\<sigma>\<close> from \<open>F\<close> to \<open>G\<close> and a pseudonatural transformation \<open>\<rho>\<close>
from \<open>G\<close> to \<open>H\<close> can be composed to obtain a pseudonatural transformation \<open>\<tau>\<close> from
\<open>F\<close> to \<open>H\<close>. The components at objects are composed via horizontal composition.
Composing the components at 1-cells is straightforward, but is formally complicated by
the need for associativities. The additional complexity leads to somewhat lengthy
proofs of the coherence conditions. This issue only gets worse as we consider additional
constructions on pseudonatural transformations.
\<close>
locale composite_pseudonatural_transformation =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F +
G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G +
H: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H +
\<sigma>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<sigma>\<^sub>0 \<sigma>\<^sub>1 +
\<rho>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G H \<Phi>\<^sub>H \<rho>\<^sub>0 \<rho>\<^sub>1
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and H :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>H :: "'c * 'c \<Rightarrow> 'd"
and \<sigma>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<sigma>\<^sub>1 :: "'c \<Rightarrow> 'd"
and \<rho>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<rho>\<^sub>1 :: "'c \<Rightarrow> 'd"
begin
definition map\<^sub>0
where "map\<^sub>0 a = \<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a"
definition map\<^sub>1
where "map\<^sub>1 f = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C f), \<sigma>\<^sub>0 (trg\<^sub>C f), F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]"
sublocale pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F H \<Phi>\<^sub>H map\<^sub>0 map\<^sub>1
proof
show "\<And>a. C.obj a \<Longrightarrow> D.ide (map\<^sub>0 a)"
unfolding map\<^sub>0_def by fastforce
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>map\<^sub>0 a : src\<^sub>D (F a) \<rightarrow>\<^sub>D src\<^sub>D (H a)\<guillemotright>"
unfolding map\<^sub>0_def
using \<rho>.map\<^sub>0_in_hhom \<sigma>.map\<^sub>0_in_hhom by blast
show "\<And>f. C.ide f \<Longrightarrow> \<guillemotleft>map\<^sub>1 f : H f \<star>\<^sub>D map\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F f\<guillemotright>"
by (unfold map\<^sub>1_def map\<^sub>0_def, intro D.comp_in_homI) auto
show "\<And>f. C.ide f \<Longrightarrow> D.iso (map\<^sub>1 f)"
unfolding map\<^sub>1_def
using \<rho>.ide_map\<^sub>0_obj \<rho>.iso_map\<^sub>1_ide \<sigma>.ide_map\<^sub>0_obj \<sigma>.iso_map\<^sub>1_ide
by (intro D.isos_compose) auto
show "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> map\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (H \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>C \<mu>))
= (map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D map\<^sub>1 (C.dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "C.arr \<mu>"
have "map\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (H \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>C \<mu>))
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), \<sigma>\<^sub>0 (trg\<^sub>C \<mu>), F (C.cod \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), G (C.cod \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 (C.cod \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (C.cod \<mu>), \<rho>\<^sub>0 (src\<^sub>C \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
(H \<mu> \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))"
unfolding map\<^sub>1_def map\<^sub>0_def
using \<mu> D.comp_assoc by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), \<sigma>\<^sub>0 (trg\<^sub>C \<mu>), F (C.cod \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), G (C.cod \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
((\<rho>\<^sub>1 (C.cod \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
((H \<mu> \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C \<mu>)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (C.dom \<mu>), \<rho>\<^sub>0 (src\<^sub>C \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)]"
using \<mu> D.assoc'_naturality [of "H \<mu>" "\<rho>\<^sub>0 (src\<^sub>C \<mu>)" "\<sigma>\<^sub>0 (src\<^sub>C \<mu>)"] D.comp_assoc by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), \<sigma>\<^sub>0 (trg\<^sub>C \<mu>), F (C.cod \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), G (C.cod \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (C.dom \<mu>), \<rho>\<^sub>0 (src\<^sub>C \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)]"
proof -
have "(\<rho>\<^sub>1 (C.cod \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D ((H \<mu> \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C \<mu>)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) =
\<rho>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (H \<mu> \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C \<mu>)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)"
using \<mu> D.whisker_right by simp
also have "... = (\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D \<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)"
using \<mu> \<rho>.naturality by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))"
using \<mu> D.whisker_right by simp
finally have
"(\<rho>\<^sub>1 (C.cod \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D ((H \<mu> \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C \<mu>)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) =
((\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D (\<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), \<sigma>\<^sub>0 (trg\<^sub>C \<mu>), F (C.cod \<mu>)] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), G (C.dom \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (C.dom \<mu>), \<rho>\<^sub>0 (src\<^sub>C \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)]"
using \<mu> D.assoc_naturality [of "\<rho>\<^sub>0 (trg\<^sub>C \<mu>)" "G \<mu>" "\<sigma>\<^sub>0 (src\<^sub>C \<mu>)"] D.comp_assoc
by simp
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), \<sigma>\<^sub>0 (trg\<^sub>C \<mu>), F (C.cod \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.dom \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), G (C.dom \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (C.dom \<mu>), \<rho>\<^sub>0 (src\<^sub>C \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)]"
proof -
have "(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) =
\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (G \<mu> \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>))"
using \<mu> D.whisker_left by simp
also have "... = \<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D (\<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<sigma>\<^sub>1 (C.dom \<mu>)"
using \<mu> \<sigma>.naturality by simp
also have "... = (\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.dom \<mu>))"
using \<mu> D.whisker_left by simp
finally have "(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.cod \<mu>)) \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) =
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.dom \<mu>))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C \<mu>)) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), \<sigma>\<^sub>0 (trg\<^sub>C \<mu>), F (C.dom \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D \<sigma>\<^sub>1 (C.dom \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C \<mu>), G (C.dom \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 (C.dom \<mu>) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (C.dom \<mu>), \<rho>\<^sub>0 (src\<^sub>C \<mu>), \<sigma>\<^sub>0 (src\<^sub>C \<mu>)]"
using \<mu> D.assoc'_naturality [of "\<rho>\<^sub>0 (trg\<^sub>C \<mu>)" "\<sigma>\<^sub>0 (trg\<^sub>C \<mu>)" "F \<mu>"] D.comp_assoc
by simp
also have "... = (map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D map\<^sub>1 (C.dom \<mu>)"
unfolding map\<^sub>1_def map\<^sub>0_def
using \<mu> by simp
finally show "map\<^sub>1 (C.cod \<mu>) \<cdot>\<^sub>D (H \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>C \<mu>)) =
(map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D map\<^sub>1 (C.dom \<mu>)"
by blast
qed
show "\<And>a. C.obj a \<Longrightarrow> (map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]
= map\<^sub>1 a \<cdot>\<^sub>D (H.unit a \<star>\<^sub>D map\<^sub>0 a)"
proof -
fix a
assume a: "C.obj a"
have "map\<^sub>1 a \<cdot>\<^sub>D (H.unit a \<star>\<^sub>D map\<^sub>0 a)
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F a] \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(\<rho>\<^sub>1 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(H.unit a \<star>\<^sub>D \<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a)"
unfolding map\<^sub>1_def map\<^sub>0_def
using a C.obj_simps D.comp_assoc by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F a] \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
((\<rho>\<^sub>1 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
((H.unit a \<star>\<^sub>D \<rho>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a]"
using a C.obj_simps D.assoc'_naturality [of "H.unit a" "\<rho>\<^sub>0 a" "\<sigma>\<^sub>0 a"] D.comp_assoc
by auto
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F a] \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 a, G a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
((\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a) \<star>\<^sub>D \<sigma>\<^sub>0 a)) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a]"
proof -
have "(\<rho>\<^sub>1 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D ((H.unit a \<star>\<^sub>D \<rho>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a)
= \<rho>\<^sub>1 a \<cdot>\<^sub>D (H.unit a \<star>\<^sub>D \<rho>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a"
using a C.obj_simps D.whisker_right
by (metis C.objE D.hcomp_simps(4) D.hseqI' D.ideD(1) D.ideD(3) D.seqI
H.unit_simps(1) H.unit_simps(2) H.unit_simps(5)
\<rho>.ide_map\<^sub>0_obj \<rho>.map\<^sub>0_simps(3) \<rho>.map\<^sub>1_simps(1) \<rho>.map\<^sub>1_simps(4) \<sigma>.ide_map\<^sub>0_obj)
also have "... = (\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a"
using a C.obj_simps \<rho>.respects_unit D.comp_assoc by simp
also have "... = ((\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a) \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D (\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a)"
using a C.obj_simps D.whisker_right
by (metis C.objE D.hcomp_simps(4) D.hseqI' D.ideD(1) D.ideD(3) D.seqI D.trg_cod
H.unit_simps(1-2,5) H.\<i>_simps(3,5) \<rho>.ide_map\<^sub>0_obj \<rho>.map\<^sub>0_simps(3)
\<rho>.map\<^sub>1_simps(1,4) \<rho>.respects_unit \<sigma>.ide_map\<^sub>0_obj)
finally have "(\<rho>\<^sub>1 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D ((H.unit a \<star>\<^sub>D \<rho>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a) =
((\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a) \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D (\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F a] \<cdot>\<^sub>D
((\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a) \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<sigma>\<^sub>0 a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a]"
using a D.assoc_naturality [of "\<rho>\<^sub>0 a" "G.unit a" "\<sigma>\<^sub>0 a"] D.comp_assoc by auto
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F a] \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a \<star>\<^sub>D F.unit a)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<sigma>\<^sub>0 a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a]"
proof -
have "(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a) \<cdot>\<^sub>D (\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<sigma>\<^sub>0 a)
= \<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a \<cdot>\<^sub>D (G.unit a \<star>\<^sub>D \<sigma>\<^sub>0 a)"
using a D.whisker_left [of "\<rho>\<^sub>0 a" "\<sigma>\<^sub>1 a" "G.unit a \<star>\<^sub>D \<sigma>\<^sub>0 a"] by force
also have "... = \<rho>\<^sub>0 a \<star>\<^sub>D (\<sigma>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<sigma>\<^sub>0 a]"
using a \<sigma>.respects_unit by simp
also have "... = (\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D (\<rho>\<^sub>0 a \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<sigma>\<^sub>0 a])"
using a D.whisker_left C.obj_simps by fastforce
finally have "(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>1 a) \<cdot>\<^sub>D (\<rho>\<^sub>0 a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<sigma>\<^sub>0 a) =
(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D (\<rho>\<^sub>0 a \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<sigma>\<^sub>0 a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F.map\<^sub>0 a] \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<sigma>\<^sub>0 a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a]"
using a D.assoc'_naturality [of "\<rho>\<^sub>0 a" "\<sigma>\<^sub>0 a" "F.unit a"] D.comp_assoc
by fastforce
also have "... = ((\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, \<sigma>\<^sub>0 a, F.map\<^sub>0 a] \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 a])) \<cdot>\<^sub>D
(\<rho>\<^sub>0 a \<star>\<^sub>D \<l>\<^sub>D[\<sigma>\<^sub>0 a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
(\<l>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \<rho>\<^sub>0 a, \<sigma>\<^sub>0 a]"
using a D.whisker_left D.whisker_right D.comp_assoc by simp
also have "... = ((\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
(\<r>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a)) \<cdot>\<^sub>D
\<l>\<^sub>D[\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a]"
using a D.lunit_hcomp(3) [of "\<rho>\<^sub>0 a" "\<sigma>\<^sub>0 a"] D.runit_hcomp(2) [of "\<rho>\<^sub>0 a" "\<sigma>\<^sub>0 a"]
D.triangle' [of "\<rho>\<^sub>0 a" "\<sigma>\<^sub>0 a"] D.comp_assoc
by auto
also have "... = ((\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a] \<cdot>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 a, G.map\<^sub>0 a, \<sigma>\<^sub>0 a]
= (\<rho>\<^sub>0 a \<star>\<^sub>D G.map\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a"
using a D.comp_inv_arr'
by (metis C.obj_def' D.comp_assoc_assoc'(2) G.map\<^sub>0_def G.map\<^sub>0_simps(1)
G.preserves_trg G.weak_arrow_of_homs_axioms \<rho>.ide_map\<^sub>0_obj \<rho>.map\<^sub>0_simps(2)
\<sigma>.ide_map\<^sub>0_obj \<sigma>.map\<^sub>0_simps(3) horizontal_homs.objE weak_arrow_of_homs_def)
moreover have "((\<rho>\<^sub>0 a \<star>\<^sub>D G.map\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D (\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a)
= (\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a)"
using a D.comp_cod_arr D.whisker_right
by (metis D.runit'_simps(1) D.runit'_simps(5) G.map\<^sub>0_def \<rho>.ide_map\<^sub>0_obj
\<rho>.map\<^sub>0_simps(2) \<sigma>.ide_map\<^sub>0_obj)
moreover have "(\<r>\<^sub>D[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D (\<r>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 a] \<star>\<^sub>D \<sigma>\<^sub>0 a) = \<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a"
using a D.whisker_right D.comp_arr_inv' D.R.as_nat_iso.components_are_iso
by (metis D.ideD(1) D.iso_runit D.runit_simps(5) \<rho>.ide_map\<^sub>0_obj \<sigma>.ide_map\<^sub>0_obj)
moreover have "(\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a) \<cdot>\<^sub>D \<l>\<^sub>D[\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a] = \<l>\<^sub>D[\<rho>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a]"
using a D.comp_cod_arr \<open>\<And>a. C.obj a \<Longrightarrow> D.ide (map\<^sub>0 a)\<close> map\<^sub>0_def by auto
ultimately show ?thesis
using D.comp_assoc by metis
qed
also have "... = (map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]"
unfolding map\<^sub>0_def by simp
finally show "(map\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]
= map\<^sub>1 a \<cdot>\<^sub>D (H.unit a \<star>\<^sub>D map\<^sub>0 a)"
by simp
qed
show "\<And>f g. \<lbrakk>C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\<rbrakk> \<Longrightarrow>
(map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[H g, map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(H g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[H g, H f, map\<^sub>0 (src\<^sub>C f)]
= map\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>C f))"
proof -
fix f g
assume f: "C.ide f" and g: "C.ide g" and fg: "src\<^sub>C g = trg\<^sub>C f"
have "(map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[H g, map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(H g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[H g, H f, map\<^sub>0 (src\<^sub>C f)]
= ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g)]
\<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C f), \<sigma>\<^sub>0 (trg\<^sub>C f), F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
unfolding map\<^sub>0_def map\<^sub>1_def
using f g fg by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C f), \<sigma>\<^sub>0 (trg\<^sub>C f), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.whisker_left D.whisker_right D.comp_assoc by fastforce
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
(((\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C f), \<sigma>\<^sub>0 (trg\<^sub>C f), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.pentagon' D.comp_assoc
D.invert_side_of_triangle(2)
[of "\<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f]"
"(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g), F f]"
"H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]"]
by force
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f] \<cdot>\<^sub>D
(((H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C f), \<sigma>\<^sub>0 (trg\<^sub>C f), F f])) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "((\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f)"
using f g fg D.assoc'_naturality [of "\<rho>\<^sub>1 g" "\<sigma>\<^sub>0 (src\<^sub>C g)" "F f"] by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "((H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C f), \<sigma>\<^sub>0 (trg\<^sub>C f), F f])) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f)
= H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f"
proof -
have "(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f])
= H g \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f"
proof -
have "(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f])
= H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (src\<^sub>C g), \<sigma>\<^sub>0 (src\<^sub>C g), F f]"
using f g fg D.whisker_left by simp
also have "... = H g \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f"
using f g fg D.comp_arr_inv' by simp
finally show ?thesis by simp
qed
moreover have "(H g \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f)
= (H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>1 f)"
using f g fg D.comp_cod_arr by simp
ultimately show ?thesis
using fg by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
((H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.assoc'_naturality [of "H g" "\<rho>\<^sub>0 (src\<^sub>C g)" "\<sigma>\<^sub>1 f"] D.comp_assoc by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
(\<rho>\<^sub>1 g \<star>\<^sub>D G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)])) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<rho>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "(\<rho>\<^sub>1 g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f) \<cdot>\<^sub>D ((H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>1 f) =
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D (\<rho>\<^sub>1 g \<star>\<^sub>D G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.interchange D.comp_arr_dom D.comp_cod_arr
by (metis C.ideD(1) \<rho>.map\<^sub>1_simps(1,5) \<rho>.naturality \<sigma>.map\<^sub>1_simps(1,4)
\<sigma>.naturality C.ideD(2,3))
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
(\<rho>\<^sub>1 g \<star>\<^sub>D G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)])) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
((H g \<star>\<^sub>D \<rho>\<^sub>1 f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.comp_assoc D.hcomp_reassoc(2) by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
((\<rho>\<^sub>1 g \<star>\<^sub>D G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((H g \<star>\<^sub>D \<rho>\<^sub>1 f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)])) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]
= \<a>\<^sub>D[H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
proof -
have "((H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
(\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
= \<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.pentagon D.comp_assoc by simp
moreover have "D.seq \<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]
\<a>\<^sub>D[H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg
by (intro D.seqI) auto
ultimately show ?thesis
using f g fg D.comp_assoc
D.invert_opposite_sides_of_square
[of "\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<a>\<^sub>D[H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"(H g \<star>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<a>\<^sub>D[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"]
by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((H g \<star>\<^sub>D \<rho>\<^sub>1 f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "(\<rho>\<^sub>1 g \<star>\<^sub>D G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D \<a>\<^sub>D[H g \<star>\<^sub>D \<rho>\<^sub>0 (trg\<^sub>C f), G f, \<sigma>\<^sub>0 (src\<^sub>C f)] =
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D ((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.assoc_naturality [of "\<rho>\<^sub>1 g" "G f" "\<sigma>\<^sub>0 (src\<^sub>C f)"] by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)])"
proof -
have "((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((H g \<star>\<^sub>D \<rho>\<^sub>1 f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) =
(\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
using f g fg D.whisker_right by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D
\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]
\<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
proof -
have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D
((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]
= \<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f))"
using f g fg \<rho>.respects_hcomp D.comp_assoc by simp
moreover have "D.seq (\<rho>\<^sub>1 (g \<star>\<^sub>C f)) (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char by auto
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately
have "(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D
((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f))
= (\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.invert_side_of_triangle(2) by simp
moreover have "D.seq (\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)))
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by auto
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def by auto
ultimately
have "\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D
((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f))
= (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]"
using f g fg
D.invert_side_of_triangle(1)
[of "(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]"
"\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D ((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f))"]
by simp
hence "(\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f)
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
D.invert_side_of_triangle(1)
by simp
hence "(\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (src\<^sub>C g), G f] \<cdot>\<^sub>D (H g \<star>\<^sub>D \<rho>\<^sub>1 f)
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D
\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)]"
using D.comp_assoc by simp
thus ?thesis
using fg by simp
qed
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
- using f g fg D.whisker_right C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg D.whisker_right C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
H.FF_def G.FF_def
by force
finally have "((\<rho>\<^sub>1 g \<star>\<^sub>D G f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, \<rho>\<^sub>0 (trg\<^sub>C f), G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((H g \<star>\<^sub>D \<rho>\<^sub>1 f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
= (\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "((\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)])) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]
= \<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)])
= \<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"
using f g fg D.pentagon' D.comp_assoc by simp
moreover have "D.seq (\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]))"
using f g fg by auto
ultimately show ?thesis
using f g fg D.comp_assoc
D.invert_side_of_triangle(2)
[of "(\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f)] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(H g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)])"
"\<a>\<^sub>D\<^sup>-\<^sup>1[H g \<star>\<^sub>D H f, \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<a>\<^sub>D\<^sup>-\<^sup>1[H g, H f, \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)]"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.assoc'_naturality [of "\<Phi>\<^sub>H (g, f)" "\<rho>\<^sub>0 (src\<^sub>C f)" "\<sigma>\<^sub>0 (src\<^sub>C f)"]
by fastforce
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F g, F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.hcomp_reassoc(1) D.comp_assoc by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F g, F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g) \<star>\<^sub>D \<sigma>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.comp_assoc D.pentagon
D.invert_opposite_sides_of_square
[of "\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g)] \<star>\<^sub>D F f)"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, \<sigma>\<^sub>0 (src\<^sub>C g) \<star>\<^sub>D F f]"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]"]
by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F g, F f] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g \<star>\<^sub>D \<sigma>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.assoc_naturality [of "\<rho>\<^sub>0 (trg\<^sub>C g)" "G g" "\<sigma>\<^sub>1 f"] D.comp_assoc by simp
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F g, F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f])) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
proof -
have "(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g \<star>\<^sub>D \<sigma>\<^sub>1 f)
= \<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D (\<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<sigma>\<^sub>1 f)"
using f g fg D.whisker_left by simp
also have "... = \<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
\<sigma>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
proof -
have "((\<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<sigma>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]
= \<sigma>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg \<sigma>.respects_hcomp D.comp_assoc by simp
hence "(\<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<sigma>\<^sub>1 f)
= (\<sigma>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char D.invert_side_of_triangle(2)
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.invert_side_of_triangle(2)
by simp
hence "\<a>\<^sub>D[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(\<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<sigma>\<^sub>1 f)
= (\<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(\<sigma>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char G.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C G.FF_def
D.invert_side_of_triangle(1)
[of "(\<sigma>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)"
"\<a>\<^sub>D[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D (\<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<sigma>\<^sub>1 f)"]
by simp
hence "(\<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<sigma>\<^sub>1 f)
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D (\<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(\<sigma>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char G.FF_def F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C G.FF_def F.FF_def
D.invert_side_of_triangle(1)
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f]) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)])"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def G.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def G.FF_def
D.whisker_left
by force
finally have "(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<sigma>\<^sub>0 (src\<^sub>C g), F f]) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g \<star>\<^sub>D \<sigma>\<^sub>1 f)
= (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f]) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g \<star>\<^sub>D F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f)))) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.pentagon' D.comp_assoc
D.invert_side_of_triangle(1)
[of "(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g] \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D F g, F f] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f])"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g), F g, F f]"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F g \<star>\<^sub>D F f]" ]
by simp
also have "... = ((((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f)))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.assoc'_naturality [of "\<rho>\<^sub>0 (trg\<^sub>C g)" "\<sigma>\<^sub>0 (trg\<^sub>C g)" "D.inv (\<Phi>\<^sub>F (g, f))"]
D.comp_assoc
by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
proof -
have "(((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f)))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)]"
proof -
have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f)))
= (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f))"
- using f g fg D.whisker_left C.VV.arr_char by simp
+ using f g fg D.whisker_left C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D F (g \<star>\<^sub>C f)"
using f g fg D.comp_arr_inv' F.cmp_components_are_iso by simp
finally have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D D.inv (\<Phi>\<^sub>F (g, f)))
= (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D F (g \<star>\<^sub>C f)"
by blast
moreover have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C g)) \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)]"
using f g fg D.comp_cod_arr by simp
ultimately show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.pentagon D.comp_assoc
D.invert_side_of_triangle(1)
[of "\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"]
D.invert_side_of_triangle(2)
[of "(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G g, G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]"
"\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g, G f] \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"]
by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G (g \<star>\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
((((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
proof -
have "(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G g \<star>\<^sub>D G f, \<sigma>\<^sub>0 (src\<^sub>C f)]
= \<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G (g \<star>\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
using f g fg D.assoc_naturality [of "\<rho>\<^sub>0 (trg\<^sub>C g)" "\<Phi>\<^sub>G (g, f)" "\<sigma>\<^sub>0 (src\<^sub>C f)"]
by fastforce
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<rho>\<^sub>0 (trg\<^sub>C g), \<sigma>\<^sub>0 (trg\<^sub>C g), F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<sigma>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<rho>\<^sub>0 (trg\<^sub>C g), G (g \<star>\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (g \<star>\<^sub>C f), \<rho>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<Phi>\<^sub>H (g, f) \<star>\<^sub>D \<rho>\<^sub>0 (src\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))"
proof -
have "(((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
= \<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
proof -
have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
= (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
(\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)))
\<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
- using f g fg C.VV.arr_char D.whisker_right by simp
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.whisker_right by simp
also have "... = (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
- using f g fg C.VV.arr_char D.whisker_left [of "\<rho>\<^sub>0 (trg\<^sub>C g)"] by simp
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.whisker_left [of "\<rho>\<^sub>0 (trg\<^sub>C g)"] by simp
also have "... = (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G (g \<star>\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
using f g fg D.comp_arr_inv' G.cmp_components_are_iso G.cmp_simps(5) by auto
finally have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
= (\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G (g \<star>\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
by blast
moreover have "((\<rho>\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D G (g \<star>\<^sub>C f)) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
(\<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f))
= \<rho>\<^sub>1 (g \<star>\<^sub>C f) \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)"
using f g fg D.comp_cod_arr by simp
ultimately show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = map\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>C f))"
unfolding map\<^sub>0_def map\<^sub>1_def
using f g fg D.comp_assoc by simp
finally show "(map\<^sub>0 (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>C g), F g, F f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[H g, map\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D
(H g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[H g, H f, map\<^sub>0 (src\<^sub>C f)]
= map\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>H (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>C f))"
using D.comp_assoc by simp
qed
qed
lemma is_pseudonatural_transformation:
shows "pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F H \<Phi>\<^sub>H map\<^sub>0 map\<^sub>1"
..
end
subsection "Whiskering of Pseudonatural Transformations"
text \<open>
Similarly to ordinary natural transformations, pseudonatural transformations can be whiskered
with pseudofunctors on the left and the right.
\<close>
locale pseudonatural_transformation_whisker_right =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
\<tau>.F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F +
\<tau>.G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G +
H: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C H \<Phi>\<^sub>H +
\<tau>: pseudonatural_transformation
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
and V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and H :: "'b \<Rightarrow> 'c"
and \<Phi>\<^sub>H :: "'b * 'b \<Rightarrow> 'c"
and \<tau>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1 :: "'c \<Rightarrow> 'd"
begin
interpretation FoH: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H F \<Phi>\<^sub>F
..
interpretation GoH: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H G \<Phi>\<^sub>G
..
definition map\<^sub>0
where "map\<^sub>0 a = \<tau>\<^sub>0 (H.map\<^sub>0 a)"
definition map\<^sub>1
where "map\<^sub>1 f = \<tau>\<^sub>1 (H f)"
sublocale pseudonatural_transformation V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>F o H\<close> FoH.cmp \<open>G o H\<close> GoH.cmp map\<^sub>0 map\<^sub>1
proof
show "\<And>a. B.obj a \<Longrightarrow> D.ide (map\<^sub>0 a)"
using map\<^sub>0_def by simp
show "\<And>a. B.obj a \<Longrightarrow> \<guillemotleft>map\<^sub>0 a : src\<^sub>D ((F \<circ> H) a) \<rightarrow>\<^sub>D src\<^sub>D ((G \<circ> H) a)\<guillemotright>"
using map\<^sub>0_def \<tau>.map\<^sub>0_in_hhom B.obj_simps C.obj_simps by simp
show "\<And>f. B.ide f \<Longrightarrow> D.iso (map\<^sub>1 f)"
using map\<^sub>1_def \<tau>.iso_map\<^sub>1_ide by simp
show "\<And>f. B.ide f \<Longrightarrow>
\<guillemotleft>map\<^sub>1 f : (G \<circ> H) f \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f) \<Rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>B f) \<star>\<^sub>D (F \<circ> H) f\<guillemotright>"
using map\<^sub>0_def map\<^sub>1_def by auto
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow>
map\<^sub>1 (B.cod \<mu>) \<cdot>\<^sub>D ((G \<circ> H) \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>B \<mu>)) =
(map\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>D (F \<circ> H) \<mu>) \<cdot>\<^sub>D map\<^sub>1 (B.dom \<mu>)"
unfolding map\<^sub>0_def map\<^sub>1_def
using \<tau>.naturality by force
show "\<And>a. B.obj a \<Longrightarrow>
(map\<^sub>0 a \<star>\<^sub>D FoH.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]
= map\<^sub>1 a \<cdot>\<^sub>D (GoH.unit a \<star>\<^sub>D map\<^sub>0 a)"
proof -
fix a
assume a: "B.obj a"
have "map\<^sub>1 a \<cdot>\<^sub>D (GoH.unit a \<star>\<^sub>D map\<^sub>0 a) =
\<tau>\<^sub>1 (H a) \<cdot>\<^sub>D (G (H.unit a) \<cdot>\<^sub>D \<tau>.G.unit (H.map\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0 (H.map\<^sub>0 a))"
unfolding map\<^sub>0_def map\<^sub>1_def
using a GoH.unit_char' by simp
also have "... = (\<tau>\<^sub>1 (H a) \<cdot>\<^sub>D (G (H.unit a) \<star>\<^sub>D \<tau>\<^sub>0 (H.map\<^sub>0 a))) \<cdot>\<^sub>D
(\<tau>.G.unit (H.map\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0 (H.map\<^sub>0 a))"
using a D.whisker_right D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0 (H.map\<^sub>0 a) \<star>\<^sub>D F (H.unit a)) \<cdot>\<^sub>D \<tau>\<^sub>1 (H.map\<^sub>0 a) \<cdot>\<^sub>D
(\<tau>.G.unit (H.map\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0 (H.map\<^sub>0 a))"
using a \<tau>.naturality [of "H.unit a"] D.comp_assoc by simp
also have "... = ((\<tau>\<^sub>0 (H.map\<^sub>0 a) \<star>\<^sub>D F (H.unit a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (H.map\<^sub>0 a) \<star>\<^sub>D \<tau>.F.unit (H.map\<^sub>0 a))) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 (H.map\<^sub>0 a)] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 (H.map\<^sub>0 a)]"
using a \<tau>.respects_unit D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0 (H.map\<^sub>0 a) \<star>\<^sub>D F (H.unit a) \<cdot>\<^sub>D \<tau>.F.unit (H.map\<^sub>0 a)) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 (H.map\<^sub>0 a)] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 (H.map\<^sub>0 a)]"
using a D.whisker_left by simp
also have "... = (map\<^sub>0 a \<star>\<^sub>D FoH.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]"
unfolding map\<^sub>0_def map\<^sub>1_def
using a FoH.unit_char' by simp
finally show "(map\<^sub>0 a \<star>\<^sub>D FoH.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]
= map\<^sub>1 a \<cdot>\<^sub>D (GoH.unit a \<star>\<^sub>D map\<^sub>0 a)"
by simp
qed
show "\<And>f g. \<lbrakk>B.ide f; B.ide g; src\<^sub>B g = trg\<^sub>B f\<rbrakk> \<Longrightarrow>
(map\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>D FoH.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (F \<circ> H) g, (F \<circ> H) f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D (F \<circ> H) f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[(G \<circ> H) g, map\<^sub>0 (src\<^sub>B g), (F \<circ> H) f] \<cdot>\<^sub>D
((G \<circ> H) g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[(G \<circ> H) g, (G \<circ> H) f, map\<^sub>0 (src\<^sub>B f)]
= map\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>D (GoH.cmp (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f))"
proof -
fix f g
assume f: "B.ide f" and g: "B.ide g" and fg: "src\<^sub>B g = trg\<^sub>B f"
have "map\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>D (GoH.cmp (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f))
= \<tau>\<^sub>1 (H (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
(G (H (g \<star>\<^sub>B f)) \<cdot>\<^sub>D G (\<Phi>\<^sub>H (g, f)) \<cdot>\<^sub>D \<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
unfolding map\<^sub>0_def map\<^sub>1_def
- using f g fg GoH.cmp_def B.VV.arr_char B.VV.dom_char by simp
+ using f g fg GoH.cmp_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = (\<tau>\<^sub>1 (H (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
(G (H (g \<star>\<^sub>B f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>H (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char H.FF_def
- C.VV.arr_char C.VV.dom_char C.VV.cod_char D.whisker_right
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.whisker_right
D.comp_assoc
by simp
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (H (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>H (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
using f g fg \<tau>.naturality [of "H (g \<star>\<^sub>B f)"] D.comp_assoc by auto
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (H (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (\<Phi>\<^sub>H (g, f))) \<cdot>\<^sub>D
\<tau>\<^sub>1 (H g \<star>\<^sub>C H f) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
using f g fg \<tau>.naturality [of "\<Phi>\<^sub>H (g, f)"] D.comp_assoc by fastforce
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (H (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (\<Phi>\<^sub>H (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D \<Phi>\<^sub>F (H g, H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))))"
proof -
have "\<tau>\<^sub>1 (H g \<star>\<^sub>C H f) = (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D \<Phi>\<^sub>F (H g, H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))] \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
proof -
have "\<tau>\<^sub>1 (H g \<star>\<^sub>C H f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))
= (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D \<Phi>\<^sub>F (H g, H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))]"
using f g fg \<tau>.respects_hcomp [of "H f" "H g"] by simp
moreover have "D.seq (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D \<Phi>\<^sub>F (H g, H f))
(\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))])"
- using f g fg C.obj_simps C.VV.arr_char C.VV.dom_char C.VV.cod_char \<tau>.F.FF_def
+ using f g fg C.obj_simps C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<tau>.F.FF_def
by simp
moreover have "D.iso (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
using f g fg \<tau>.G.cmp_components_are_iso [of "H g" "H f"] by simp
ultimately
have "\<tau>\<^sub>1 (H g \<star>\<^sub>C H f) = ((\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D \<Phi>\<^sub>F (H g, H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))]) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
using D.invert_side_of_triangle(2) by blast
thus ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (H (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (\<Phi>\<^sub>H (g, f))) \<cdot>\<^sub>D
(\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D \<Phi>\<^sub>F (H g, H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))]"
proof -
have "(D.inv (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))))
= (G (H g) \<star>\<^sub>D G (H f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))"
proof -
have "(D.inv (\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))))
= (D.inv (\<Phi>\<^sub>G (H g, H f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f)))"
using f g fg by simp
also have "... = D.inv (\<Phi>\<^sub>G (H g, H f)) \<cdot>\<^sub>D \<Phi>\<^sub>G (H g, H f) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))"
- using f g fg D.whisker_right C.VV.arr_char by simp
+ using f g fg D.whisker_right C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = (G (H g) \<star>\<^sub>D G (H f)) \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C (H f))"
using f g fg D.comp_inv_arr' \<tau>.G.cmp_components_are_iso by simp
finally show ?thesis by blast
qed
moreover have "\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))] \<cdot>\<^sub>D ...
= \<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))]"
using f g fg D.comp_arr_dom by simp
ultimately show ?thesis by simp
qed
also have "... = (\<tau>\<^sub>0 (trg\<^sub>C (H g)) \<star>\<^sub>D F (H (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
F (\<Phi>\<^sub>H (g, f)) \<cdot>\<^sub>D
\<Phi>\<^sub>F (H g, H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 (H g) \<star>\<^sub>D F (H f)) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[G (H g), \<tau>\<^sub>0 (src\<^sub>C (H g)), F (H f)] \<cdot>\<^sub>D
(G (H g) \<star>\<^sub>D \<tau>\<^sub>1 (H f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (H g), G (H f), \<tau>\<^sub>0 (src\<^sub>C (H f))]"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char
- B.VV.arr_char B.VV.dom_char B.VV.cod_char H.FF_def
- D.whisker_left C.VV.arr_char B.VV.arr_char D.comp_assoc
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
+ D.whisker_left C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.comp_assoc
by auto
also have "... = (map\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>D FoH.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (F \<circ> H) g, (F \<circ> H) f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D (F \<circ> H) f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[(G \<circ> H) g, map\<^sub>0 (src\<^sub>B g), (F \<circ> H) f] \<cdot>\<^sub>D
((G \<circ> H) g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[(G \<circ> H) g, (G \<circ> H) f, map\<^sub>0 (src\<^sub>B f)]"
unfolding map\<^sub>0_def map\<^sub>1_def
- using f g fg FoH.cmp_def B.VV.arr_char B.VV.dom_char by simp
+ using f g fg FoH.cmp_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C by simp
finally show "(map\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>D FoH.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (F \<circ> H) g, (F \<circ> H) f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D (F \<circ> H) f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[(G \<circ> H) g, map\<^sub>0 (src\<^sub>B g), (F \<circ> H) f] \<cdot>\<^sub>D
((G \<circ> H) g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D
\<a>\<^sub>D[(G \<circ> H) g, (G \<circ> H) f, map\<^sub>0 (src\<^sub>B f)]
= map\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>D (GoH.cmp (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f))"
by simp
qed
qed
lemma is_pseudonatural_transformation:
shows "pseudonatural_transformation V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F o H) FoH.cmp (G o H) GoH.cmp map\<^sub>0 map\<^sub>1"
..
end
locale pseudonatural_transformation_whisker_left =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
\<tau>.F: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F +
\<tau>.G: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G +
H: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H +
\<tau>: pseudonatural_transformation
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
and V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'b \<Rightarrow> 'c"
and \<Phi>\<^sub>F :: "'b * 'b \<Rightarrow> 'c"
and G :: "'b \<Rightarrow> 'c"
and \<Phi>\<^sub>G :: "'b * 'b \<Rightarrow> 'c"
and H :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>H :: "'c * 'c \<Rightarrow> 'd"
and \<tau>\<^sub>0 :: "'b \<Rightarrow> 'c"
and \<tau>\<^sub>1 :: "'b \<Rightarrow> 'c"
begin
interpretation HoF: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F H \<Phi>\<^sub>H
..
interpretation HoG: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G H \<Phi>\<^sub>H
..
definition map\<^sub>0
where "map\<^sub>0 a = H (\<tau>\<^sub>0 a)"
definition map\<^sub>1
where "map\<^sub>1 f = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B f), F f)) \<cdot>\<^sub>D H (\<tau>\<^sub>1 f) \<cdot>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f))"
sublocale pseudonatural_transformation V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
HoF.map HoF.cmp HoG.map HoG.cmp map\<^sub>0 map\<^sub>1
proof
show "\<And>a. B.obj a \<Longrightarrow> D.ide (map\<^sub>0 a)"
using map\<^sub>0_def by simp
show "\<And>f. B.ide f \<Longrightarrow> D.iso (map\<^sub>1 f)"
proof -
fix f
assume f: "B.ide f"
have "D.seq (H (\<tau>\<^sub>1 f)) (\<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f)))"
using f by (intro D.seqI) auto
moreover have "D.seq (D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B f), F f)))
(H (\<tau>\<^sub>1 f) \<cdot>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f)))"
using f \<tau>.map\<^sub>1_in_hom [of f] calculation by (intro D.seqI) auto
ultimately show "D.iso (map\<^sub>1 f)"
using f map\<^sub>1_def H.preserves_iso \<tau>.iso_map\<^sub>1_ide H.cmp_components_are_iso
- C.VV.arr_char D.isos_compose
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.isos_compose
by auto
qed
show "\<And>a. B.obj a \<Longrightarrow> \<guillemotleft>map\<^sub>0 a : src\<^sub>D ((H \<circ> F) a) \<rightarrow>\<^sub>D src\<^sub>D ((H \<circ> G) a)\<guillemotright>"
using map\<^sub>0_def by fastforce
show "\<And>f. B.ide f \<Longrightarrow>
\<guillemotleft>map\<^sub>1 f : (H \<circ> G) f \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f) \<Rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>B f) \<star>\<^sub>D (H \<circ> F) f\<guillemotright>"
using map\<^sub>0_def map\<^sub>1_def by fastforce
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow>
map\<^sub>1 (B.cod \<mu>) \<cdot>\<^sub>D ((H \<circ> G) \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>B \<mu>))
= (map\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>D (H \<circ> F) \<mu>) \<cdot>\<^sub>D map\<^sub>1 (B.dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "B.arr \<mu>"
have "(map\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>D (H \<circ> F) \<mu>) \<cdot>\<^sub>D map\<^sub>1 (B.dom \<mu>)
= (H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>)) \<star>\<^sub>D (H \<circ> F) \<mu>) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B (B.dom \<mu>)), F (B.dom \<mu>))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (B.dom \<mu>)) \<cdot>\<^sub>D \<Phi>\<^sub>H (G (B.dom \<mu>), \<tau>\<^sub>0 (src\<^sub>B (B.dom \<mu>)))"
unfolding map\<^sub>0_def map\<^sub>1_def
using \<mu> by simp
also have "... = ((H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>)) \<star>\<^sub>D H (F \<mu>)) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.dom \<mu>)))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (B.dom \<mu>)) \<cdot>\<^sub>D \<Phi>\<^sub>H (G (B.dom \<mu>), \<tau>\<^sub>0 (src\<^sub>B \<mu>))"
using \<mu> D.comp_assoc by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>C F \<mu>) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (B.dom \<mu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (B.dom \<mu>), \<tau>\<^sub>0 (src\<^sub>B (B.dom \<mu>)))"
proof -
have "\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>)) \<cdot>\<^sub>D (H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>)) \<star>\<^sub>D H (F \<mu>))
= H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>C F \<mu>) \<cdot>\<^sub>D \<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.dom \<mu>))"
using \<mu> H.\<Phi>.naturality [of "(\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F \<mu>)"]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
by simp
moreover have "D.seq (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) (H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>)) \<star>\<^sub>D H (F \<mu>))"
using \<mu>
by (intro D.seqI D.hseqI') auto
moreover have "D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>)))"
using \<mu> H.cmp_components_are_iso by simp
moreover have "D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.dom \<mu>)))"
using \<mu> H.cmp_components_are_iso by simp
ultimately
have "(H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>)) \<star>\<^sub>D H (F \<mu>)) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.dom \<mu>)))
= D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) \<cdot>\<^sub>D H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>C F \<mu>)"
- using \<mu> H.cmp_components_are_iso C.VV.arr_char D.invert_opposite_sides_of_square
+ using \<mu> H.cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.invert_opposite_sides_of_square
by blast
thus ?thesis
using \<mu> D.comp_assoc by simp
qed
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) \<cdot>\<^sub>D
H ((\<tau>\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>C F \<mu>) \<cdot>\<^sub>C \<tau>\<^sub>1 (B.dom \<mu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (B.dom \<mu>), \<tau>\<^sub>0 (src\<^sub>B (B.dom \<mu>)))"
using \<mu> by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (B.cod \<mu>) \<cdot>\<^sub>C (G \<mu> \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B \<mu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (B.dom \<mu>), \<tau>\<^sub>0 (src\<^sub>B (B.dom \<mu>)))"
using \<mu> \<tau>.naturality by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (B.cod \<mu>)) \<cdot>\<^sub>D
H (G \<mu> \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B \<mu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (B.dom \<mu>), \<tau>\<^sub>0 (src\<^sub>B (B.dom \<mu>)))"
using \<mu> D.comp_assoc by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B \<mu>), F (B.cod \<mu>))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (B.cod \<mu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (B.cod \<mu>), \<tau>\<^sub>0 (src\<^sub>B (B.cod \<mu>))) \<cdot>\<^sub>D
(H (G \<mu>) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B \<mu>)))"
using \<mu> H.\<Phi>.naturality [of "(G \<mu>, \<tau>\<^sub>0 (src\<^sub>B \<mu>))"]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
by force
also have "... = map\<^sub>1 (B.cod \<mu>) \<cdot>\<^sub>D ((H \<circ> G) \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>B \<mu>))"
unfolding map\<^sub>0_def map\<^sub>1_def
using \<mu> D.comp_assoc by simp
finally show "map\<^sub>1 (B.cod \<mu>) \<cdot>\<^sub>D ((H \<circ> G) \<mu> \<star>\<^sub>D map\<^sub>0 (src\<^sub>B \<mu>))
= (map\<^sub>0 (trg\<^sub>B \<mu>) \<star>\<^sub>D (H \<circ> F) \<mu>) \<cdot>\<^sub>D map\<^sub>1 (B.dom \<mu>)"
by simp
qed
show "\<And>a. B.obj a \<Longrightarrow> (map\<^sub>0 a \<star>\<^sub>D HoF.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]
= map\<^sub>1 a \<cdot>\<^sub>D (HoG.unit a \<star>\<^sub>D map\<^sub>0 a)"
proof -
fix a
assume a: "B.obj a"
have "map\<^sub>1 a \<cdot>\<^sub>D (HoG.unit a \<star>\<^sub>D map\<^sub>0 a)
= D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 a) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H (\<tau>.G.unit a) \<cdot>\<^sub>D H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
unfolding map\<^sub>0_def map\<^sub>1_def
using a HoG.unit_char' D.comp_assoc by auto
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 a) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H (\<tau>.G.unit a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
using a D.whisker_right D.comp_assoc by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 a) \<cdot>\<^sub>D
H (\<tau>.G.unit a \<star>\<^sub>C \<tau>\<^sub>0 a)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
using a H.\<Phi>.naturality [of "(\<tau>.G.unit a, \<tau>\<^sub>0 a)"] D.comp_assoc
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
by auto
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 a \<cdot>\<^sub>C (\<tau>.G.unit a \<star>\<^sub>C \<tau>\<^sub>0 a)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
proof -
have "C.arr (\<tau>\<^sub>1 a \<cdot>\<^sub>C (\<tau>.G.unit a \<star>\<^sub>C \<tau>\<^sub>0 a))"
using a by force
hence "H (\<tau>\<^sub>1 a) \<cdot>\<^sub>D H (\<tau>.G.unit a \<star>\<^sub>C \<tau>\<^sub>0 a) = H (\<tau>\<^sub>1 a \<cdot>\<^sub>C (\<tau>.G.unit a \<star>\<^sub>C \<tau>\<^sub>0 a))"
using a by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D
H ((\<tau>\<^sub>0 a \<star>\<^sub>C \<tau>.F.unit a) \<cdot>\<^sub>C \<r>\<^sub>C\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>C \<l>\<^sub>C[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
using a \<tau>.respects_unit by simp
also have "... = (D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D
H (\<tau>\<^sub>0 a \<star>\<^sub>C \<tau>.F.unit a)) \<cdot>\<^sub>D
H (\<r>\<^sub>C\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
H (\<l>\<^sub>C[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
using a D.comp_assoc B.obj_simps by simp
also have "... = (H (\<tau>\<^sub>0 a) \<star>\<^sub>D H (\<tau>.F.unit a)) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, \<tau>.F.map\<^sub>0 a)) \<cdot>\<^sub>D
H (\<r>\<^sub>C\<^sup>-\<^sup>1[\<tau>\<^sub>0 a])) \<cdot>\<^sub>D
H (\<l>\<^sub>C[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<cdot>\<^sub>D H (\<tau>\<^sub>0 a \<star>\<^sub>C \<tau>.F.unit a)
= (H (\<tau>\<^sub>0 a) \<star>\<^sub>D H (\<tau>.F.unit a)) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, \<tau>.F.map\<^sub>0 a))"
proof -
have "D.seq (H (\<tau>\<^sub>0 a \<star>\<^sub>C \<tau>.F.unit a)) (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, \<tau>.F.map\<^sub>0 a))"
- using a C.VV.arr_char C.VV.dom_char C.VV.cod_char C.obj_simps by auto
+ using a C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C C.obj_simps by auto
moreover have "D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)) \<and> D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, \<tau>.F.map\<^sub>0 a))"
proof -
have "C.ide (F a) \<and> src\<^sub>C (\<tau>\<^sub>0 a) = trg\<^sub>C (F a) \<and> src\<^sub>C (\<tau>\<^sub>0 a) = trg\<^sub>C (\<tau>.F.map\<^sub>0 a)"
using a by auto
moreover have "C.ide (\<tau>.F.map\<^sub>0 a)"
proof -
(* TODO: I still haven't been able to configure the simps to do this. *)
have "C.obj (\<tau>.F.map\<^sub>0 a)"
using a by simp
thus ?thesis by auto
qed
ultimately show ?thesis
using a H.cmp_components_are_iso B.obj_simps by auto
qed
ultimately show ?thesis
using a H.\<Phi>.naturality [of "(\<tau>\<^sub>0 a, \<tau>.F.unit a)"]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
D.invert_opposite_sides_of_square
[of "\<Phi>\<^sub>H (\<tau>\<^sub>0 a, F a)" "H (\<tau>\<^sub>0 a) \<star>\<^sub>D H (\<tau>.F.unit a)"
"H (\<tau>\<^sub>0 a \<star>\<^sub>C \<tau>.F.unit a)" "\<Phi>\<^sub>H (\<tau>\<^sub>0 a, \<tau>.F.map\<^sub>0 a)"]
by auto
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 a) \<star>\<^sub>D H (\<tau>.F.unit a)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 a) \<star>\<^sub>D H.unit (\<tau>.F.map\<^sub>0 a)) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D
(H (\<l>\<^sub>C[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 a, \<tau>.F.map\<^sub>0 a)) \<cdot>\<^sub>D H (\<r>\<^sub>C\<^sup>-\<^sup>1[\<tau>\<^sub>0 a])
= (H (\<tau>\<^sub>0 a) \<star>\<^sub>D H.unit (\<tau>.F.map\<^sub>0 a)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[H (\<tau>\<^sub>0 a)]"
using a H.preserves_runit(2) [of "\<tau>\<^sub>0 a"] D.comp_assoc
by (metis C.ideD(1) C.runit'_simps(1) C.src.preserves_ide C.trg_src
D.invert_side_of_triangle(1) \<tau>.F.map\<^sub>0_def H.cmp_components_are_iso
H.preserves_reflects_arr \<tau>.ide_map\<^sub>0_obj \<tau>.map\<^sub>0_simps(2))
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 a) \<star>\<^sub>D H (\<tau>.F.unit a)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 a) \<star>\<^sub>D H.unit (\<tau>.F.map\<^sub>0 a)) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D
\<l>\<^sub>D[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D
(D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a)) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
proof -
have "H (\<l>\<^sub>C[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D \<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)
= \<l>\<^sub>D[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D (D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a)) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)) \<cdot>\<^sub>D \<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)"
using a H.preserves_lunit(1) D.comp_assoc by auto
also have "... = \<l>\<^sub>D[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D (D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)) \<cdot>\<^sub>D \<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)
= H (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a)"
using a H.cmp_components_are_iso D.comp_inv_arr'
by (metis C.isomorphic_implies_hpar(1) \<tau>.G.map\<^sub>0_simps(2) \<tau>.G.weakly_preserves_objects
H.cmp_simps(4) \<tau>.ide_map\<^sub>0_obj \<tau>.map\<^sub>0_simps(3))
moreover have "(D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a)) \<cdot>\<^sub>D
(H (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))
= (D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
using a H.unit_char(2) D.comp_arr_dom
by (metis D.arr_inv D.dom_inv D.whisker_right \<tau>.G.map\<^sub>0_simps(1) H.unit_simps(5)
H.preserves_ide \<tau>.ide_map\<^sub>0_obj)
ultimately show ?thesis by simp
qed
finally have "H (\<l>\<^sub>C[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D \<Phi>\<^sub>H (\<tau>.G.map\<^sub>0 a, \<tau>\<^sub>0 a)
= \<l>\<^sub>D[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D (D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((H (\<tau>\<^sub>0 a) \<star>\<^sub>D H (\<tau>.F.unit a)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 a) \<star>\<^sub>D H.unit (\<tau>.F.map\<^sub>0 a))) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D
\<l>\<^sub>D[H (\<tau>\<^sub>0 a)]"
proof -
have "(D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a)) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))
= D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<cdot>\<^sub>D H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a)"
using a D.whisker_right H.unit_char(2) by simp
also have "... = H.map\<^sub>0 (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a)"
using a H.unit_char(1-2) [of "\<tau>.G.map\<^sub>0 a"] D.comp_inv_arr' by simp
finally have "(D.inv (H.unit (\<tau>.G.map\<^sub>0 a)) \<star>\<^sub>D H (\<tau>\<^sub>0 a)) \<cdot>\<^sub>D
(H.unit (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a))
= H.map\<^sub>0 (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a)"
by blast
moreover have "\<l>\<^sub>D[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D (H.map\<^sub>0 (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a)) = \<l>\<^sub>D[H (\<tau>\<^sub>0 a)]"
using a D.comp_arr_dom [of "\<l>\<^sub>D[H (\<tau>\<^sub>0 a)]" "H.map\<^sub>0 (\<tau>.G.map\<^sub>0 a) \<star>\<^sub>D H (\<tau>\<^sub>0 a)"]
by auto
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 a) \<star>\<^sub>D HoF.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[H (\<tau>\<^sub>0 a)] \<cdot>\<^sub>D \<l>\<^sub>D[H (\<tau>\<^sub>0 a)]"
using a HoF.unit_char' D.whisker_left [of "H (\<tau>\<^sub>0 a)"] by simp
also have "... = (map\<^sub>0 a \<star>\<^sub>D HoF.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a]"
unfolding map\<^sub>0_def by simp
finally show "(map\<^sub>0 a \<star>\<^sub>D HoF.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[map\<^sub>0 a] =
map\<^sub>1 a \<cdot>\<^sub>D (HoG.unit a \<star>\<^sub>D map\<^sub>0 a)"
by simp
qed
show "\<And>f g. \<lbrakk>B.ide f; B.ide g; src\<^sub>B g = trg\<^sub>B f\<rbrakk> \<Longrightarrow>
(map\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>D HoF.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (H \<circ> F) g, (H \<circ> F) f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D (H \<circ> F) f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[(H \<circ> G) g, map\<^sub>0 (src\<^sub>B g), (H \<circ> F) f] \<cdot>\<^sub>D
((H \<circ> G) g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[(H \<circ> G) g, (H \<circ> G) f, map\<^sub>0 (src\<^sub>B f)]
= map\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>D (HoG.cmp (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f))"
proof -
fix f g
assume f: "B.ide f" and g: "B.ide g" and fg: "src\<^sub>B g = trg\<^sub>B f"
have "map\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>D (HoG.cmp (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f))
= D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (g \<star>\<^sub>B f), \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(H (G (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
H (\<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
unfolding map\<^sub>0_def map\<^sub>1_def HoG.cmp_def
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char D.comp_assoc by simp
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.comp_assoc by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G (g \<star>\<^sub>B f), \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
((H (G (g \<star>\<^sub>B f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char
- C.VV.arr_char C.VV.dom_char C.VV.cod_char \<tau>.G.FF_def
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<tau>.G.FF_def
D.comp_assoc D.whisker_right
by auto
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G (g \<star>\<^sub>B f), \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "(H (G (g \<star>\<^sub>B f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D (H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))
= H (G (g \<star>\<^sub>B f)) \<cdot>\<^sub>D H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char D.whisker_right
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.whisker_right
by simp
also have "... = H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))"
using f g fg D.comp_cod_arr [of "H (\<Phi>\<^sub>G (g, f))" "H (G (g \<star>\<^sub>B f))"]
- B.VV.arr_char B.VV.dom_char B.VV.cod_char
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally have "(H (G (g \<star>\<^sub>B f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))
= H (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
H (\<Phi>\<^sub>G (g, f) \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
using f g fg H.\<Phi>.naturality [of "(\<Phi>\<^sub>G (g, f), \<tau>\<^sub>0 (src\<^sub>B f))"]
- B.VV.arr_char B.VV.dom_char B.VV.cod_char \<tau>.G.FF_def
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<tau>.G.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
D.comp_assoc
by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>C (\<Phi>\<^sub>G (g, f) \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char D.comp_assoc by simp
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.comp_assoc by simp
also have "... = D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
H ((\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>C
\<a>\<^sub>C[\<tau>\<^sub>0 (trg\<^sub>B g), F g, F f] \<cdot>\<^sub>C
(\<tau>\<^sub>1 g \<star>\<^sub>C F f) \<cdot>\<^sub>C
\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f] \<cdot>\<^sub>C
(G g \<star>\<^sub>C \<tau>\<^sub>1 f) \<cdot>\<^sub>C
\<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)]) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
using f g fg \<tau>.respects_hcomp by simp
also have "... = (D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D
H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C \<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
H \<a>\<^sub>C[\<tau>\<^sub>0 (trg\<^sub>B g), F g, F f] \<cdot>\<^sub>D
H (\<tau>\<^sub>1 g \<star>\<^sub>C F f) \<cdot>\<^sub>D
H (\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f]) \<cdot>\<^sub>D
H (G g \<star>\<^sub>C \<tau>\<^sub>1 f) \<cdot>\<^sub>D
H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)] \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char \<tau>.F.FF_def D.comp_assoc
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<tau>.F.FF_def D.comp_assoc
by simp
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f)) \<cdot>\<^sub>D
H \<a>\<^sub>C[\<tau>\<^sub>0 (trg\<^sub>B g), F g, F f]) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 g \<star>\<^sub>C F f) \<cdot>\<^sub>D
H (\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f]) \<cdot>\<^sub>D
H (G g \<star>\<^sub>C \<tau>\<^sub>1 f) \<cdot>\<^sub>D
H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)] \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<cdot>\<^sub>D H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C \<Phi>\<^sub>F (g, f))
= (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f))"
proof -
have "\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f)) \<cdot>\<^sub>D (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) =
H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f)"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char
- C.VV.arr_char C.VV.dom_char C.VV.cod_char \<tau>.F.FF_def H.FF_def
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<tau>.F.FF_def H.FF_def
H.\<Phi>.naturality [of "(\<tau>\<^sub>0 (trg\<^sub>B g), \<Phi>\<^sub>F (g, f))"]
by auto
moreover have "D.seq (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f)))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
by auto
moreover have "D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F (g \<star>\<^sub>B f))) \<and>
D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f))"
using f g fg H.cmp_components_are_iso by simp
ultimately show ?thesis
using f g fg D.invert_opposite_sides_of_square by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 g \<star>\<^sub>C F f)) \<cdot>\<^sub>D
H (\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f]) \<cdot>\<^sub>D
H (G g \<star>\<^sub>C \<tau>\<^sub>1 f) \<cdot>\<^sub>D
H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)] \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f)) \<cdot>\<^sub>D H \<a>\<^sub>C[\<tau>\<^sub>0 (trg\<^sub>B g), F g, F f]
= (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f)) \<cdot>\<^sub>D H \<a>\<^sub>C[\<tau>\<^sub>0 (trg\<^sub>B g), F g, F f]
= ((D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g \<star>\<^sub>C F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f))"
using f g fg H.preserves_assoc(1) D.comp_assoc by simp
also have "... = ((H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (F g \<star>\<^sub>C F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
H.FF_def D.comp_assoc D.comp_inv_arr' H.cmp_components_are_iso
by simp
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f))"
- using f g fg C.VV.arr_char H.cmp_simps(5) D.comp_cod_arr D.comp_assoc by auto
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.cmp_simps(5) D.comp_cod_arr D.comp_assoc by auto
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D
H (\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f])) \<cdot>\<^sub>D
H (G g \<star>\<^sub>C \<tau>\<^sub>1 f) \<cdot>\<^sub>D
H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)] \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f)) \<cdot>\<^sub>D H (\<tau>\<^sub>1 g \<star>\<^sub>C F f)
= (H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f))"
proof -
have "\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f) \<cdot>\<^sub>D (H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f))
= H (\<tau>\<^sub>1 g \<star>\<^sub>C F f) \<cdot>\<^sub>D \<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)"
using f g fg H.\<Phi>.naturality [of "(\<tau>\<^sub>1 g, F f)"]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
by simp
moreover have "D.seq (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f))
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f))"
using f g fg by fastforce
moreover have "D.iso (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>C F g, F f)) \<and>
D.iso (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f))"
using f g fg H.cmp_components_are_iso by simp
ultimately show ?thesis
using f g fg D.invert_opposite_sides_of_square by metis
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g) \<star>\<^sub>C F f)) \<cdot>\<^sub>D
H (G g \<star>\<^sub>C \<tau>\<^sub>1 f)) \<cdot>\<^sub>D
H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)] \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D H (\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f])
= (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g) \<star>\<^sub>C F f))"
proof -
have "D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D H (\<a>\<^sub>C\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>B g), F f])
= ((D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g) \<star>\<^sub>C F f))"
using f g fg H.preserves_assoc(2) D.comp_assoc by simp
also have "... = (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g) \<star>\<^sub>C F f))"
proof -
have "D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D \<Phi>\<^sub>H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g), F f)
= H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)"
using f g fg H.cmp_components_are_iso D.comp_inv_arr' H.FF_def
- C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "(H (G g \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f))
= (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f))"
- using f g fg D.comp_cod_arr C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg D.comp_cod_arr C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
ultimately show ?thesis by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)]) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g) \<star>\<^sub>C F f)) \<cdot>\<^sub>D H (G g \<star>\<^sub>C \<tau>\<^sub>1 f)
= (H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (trg\<^sub>B f) \<star>\<^sub>C F f) \<cdot>\<^sub>D (H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f))
= H (G g \<star>\<^sub>C \<tau>\<^sub>1 f) \<cdot>\<^sub>D \<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))"
using f g fg H.\<Phi>.naturality [of "(G g, \<tau>\<^sub>1 f)"]
- C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
by simp
thus ?thesis
using f g fg H.cmp_components_are_iso
- C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
D.invert_opposite_sides_of_square
[of "\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (trg\<^sub>B f) \<star>\<^sub>C F f)" "H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f)"
"H (G g \<star>\<^sub>C \<tau>\<^sub>1 f)" "\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g, G f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)]
= (H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g, G f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)))"
proof -
have "D.inv (\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D H \<a>\<^sub>C[G g, G f, \<tau>\<^sub>0 (src\<^sub>B f)]
= ((D.inv (\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g, G f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f)))"
using f g fg H.preserves_assoc(1)
- H.cmp_components_are_iso C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ H.cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
D.comp_assoc
by simp
moreover have "D.inv (\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g, G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))
= H (G g) \<star>\<^sub>D H (G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))"
- using f g fg H.cmp_components_are_iso C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg H.cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
D.comp_inv_arr' H.FF_def
by simp
moreover have "(H (G g) \<star>\<^sub>D H (G f \<star>\<^sub>C \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f)))
= H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
D.whisker_left D.comp_cod_arr
by simp
ultimately show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
((D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f))) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
((H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f)))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))]"
proof -
have "\<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g, G f)) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g \<star>\<^sub>C G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, G f) \<star>\<^sub>D H (\<tau>\<^sub>0 (src\<^sub>B f)))
= \<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.whisker_right [of "H (\<tau>\<^sub>0 (src\<^sub>B f))" "D.inv (\<Phi>\<^sub>H (G g, G f))" "\<Phi>\<^sub>H (G g, G f)"]
- D.comp_inv_arr' C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def
+ D.comp_inv_arr' C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C H.FF_def
H.cmp_components_are_iso
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (F (g \<star>\<^sub>B f)) \<cdot>\<^sub>D
H (\<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D
\<Phi>\<^sub>H (F g, F f)) \<cdot>\<^sub>D
\<a>\<^sub>D[H (\<tau>\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 g) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\<tau>\<^sub>0 (src\<^sub>B g)), H (F f)] \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 f) \<cdot>\<^sub>D
\<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f))) \<cdot>\<^sub>D
\<a>\<^sub>D[H (G g), H (G f), H (\<tau>\<^sub>0 (src\<^sub>B f))]"
proof -
have "(H (G g) \<star>\<^sub>D D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f))) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D H (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
(H (G g) \<star>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f)))
= H (G g) \<star>\<^sub>D
D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (src\<^sub>B g), F f)) \<cdot>\<^sub>D H (\<tau>\<^sub>1 f) \<cdot>\<^sub>D \<Phi>\<^sub>H (G f, \<tau>\<^sub>0 (src\<^sub>B f))"
- using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char D.whisker_left
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C D.whisker_left
H.cmp_components_are_iso
by simp
moreover have "(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D
(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f))
= H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D
H (F (g \<star>\<^sub>B f)) \<cdot>\<^sub>D H (\<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<Phi>\<^sub>H (F g, F f)"
proof -
have "(H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D (H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D \<Phi>\<^sub>H (F g, F f))
= H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D H (\<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<Phi>\<^sub>H (F g, F f)"
proof -
have "D.arr (H (\<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<Phi>\<^sub>H (F g, F f))"
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char \<tau>.F.FF_def
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C \<tau>.F.FF_def
by (intro D.seqI) auto
thus ?thesis
- using f g fg D.whisker_left C.VV.arr_char by simp
+ using f g fg D.whisker_left C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
qed
also have "... = H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D
(H (F (g \<star>\<^sub>B f) \<cdot>\<^sub>C \<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D \<Phi>\<^sub>H (F g, F f)"
- using f g fg C.comp_cod_arr B.VV.arr_char \<tau>.F.cmp_simps(5) by auto
+ using f g fg C.comp_cod_arr B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<tau>.F.cmp_simps(5) by auto
also have "... = H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D
(H (F (g \<star>\<^sub>B f)) \<cdot>\<^sub>D H (\<Phi>\<^sub>F (g, f))) \<cdot>\<^sub>D \<Phi>\<^sub>H (F g, F f)"
- using f g fg B.VV.arr_char \<tau>.F.cmp_simps(5) by auto
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<tau>.F.cmp_simps(5) by auto
also have "... = H (\<tau>\<^sub>0 (trg\<^sub>B g)) \<star>\<^sub>D
H (F (g \<star>\<^sub>B f)) \<cdot>\<^sub>D H (\<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<Phi>\<^sub>H (F g, F f)"
using D.comp_assoc by simp
finally show ?thesis by simp
qed
moreover have "(D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(H (\<tau>\<^sub>1 g) \<star>\<^sub>D H (F f)) \<cdot>\<^sub>D
(\<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f))
= D.inv (\<Phi>\<^sub>H (\<tau>\<^sub>0 (trg\<^sub>B g), F g)) \<cdot>\<^sub>D
H (\<tau>\<^sub>1 g) \<cdot>\<^sub>D \<Phi>\<^sub>H (G g, \<tau>\<^sub>0 (src\<^sub>B g)) \<star>\<^sub>D H (F f)"
- using f g fg D.whisker_right C.VV.arr_char C.VV.dom_char C.VV.cod_char
+ using f g fg D.whisker_right C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
ultimately show ?thesis
by simp
qed
also have "... = (map\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>D HoF.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (H \<circ> F) g, (H \<circ> F) f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D (H \<circ> F) f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[(H \<circ> G) g, map\<^sub>0 (src\<^sub>B g), (H \<circ> F) f] \<cdot>\<^sub>D
((H \<circ> G) g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[(H \<circ> G) g, (H \<circ> G) f, map\<^sub>0 (src\<^sub>B f)]"
unfolding map\<^sub>0_def map\<^sub>1_def HoF.cmp_def
- using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char by simp
+ using f g fg B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_char\<^sub>S\<^sub>b\<^sub>C by simp
finally show "(map\<^sub>0 (trg\<^sub>B g) \<star>\<^sub>D HoF.cmp (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (H \<circ> F) g, (H \<circ> F) f] \<cdot>\<^sub>D
(map\<^sub>1 g \<star>\<^sub>D (H \<circ> F) f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[(H \<circ> G) g, map\<^sub>0 (src\<^sub>B g), (H \<circ> F) f] \<cdot>\<^sub>D
((H \<circ> G) g \<star>\<^sub>D map\<^sub>1 f) \<cdot>\<^sub>D \<a>\<^sub>D[(H \<circ> G) g, (H \<circ> G) f, map\<^sub>0 (src\<^sub>B f)]
= map\<^sub>1 (g \<star>\<^sub>B f) \<cdot>\<^sub>D (HoG.cmp (g, f) \<star>\<^sub>D map\<^sub>0 (src\<^sub>B f))"
by simp
qed
qed
lemma is_pseudonatural_transformation:
shows "pseudonatural_transformation V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
HoF.map HoF.cmp HoG.map HoG.cmp map\<^sub>0 map\<^sub>1"
..
end
subsection "Pseudonatural Equivalences"
text \<open>
A \emph{pseudonatural equivalence} is a pseudonatural transformation whose components
at objects are equivalence maps. Pseudonatural equivalences between pseudofunctors
generalize natural isomorphisms between ordinary functors.
\<close>
locale pseudonatural_equivalence =
pseudonatural_transformation +
assumes components_are_equivalences: "C.obj a \<Longrightarrow> D.equivalence_map (\<tau>\<^sub>0 a)"
subsubsection "Identity Transformations are Pseudonatural Equivalences"
sublocale identity_pseudonatural_transformation \<subseteq>
pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi>\<^sub>F F \<Phi>\<^sub>F map\<^sub>0 map\<^sub>1
by unfold_locales (simp add: D.obj_is_equivalence_map)
subsubsection "Composition of Pseudonatural Equivalences"
locale composite_pseudonatural_equivalence =
composite_pseudonatural_transformation +
\<sigma>: pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<sigma>\<^sub>0 \<sigma>\<^sub>1 +
\<rho>: pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
G \<Phi>\<^sub>G H \<Phi>\<^sub>H \<rho>\<^sub>0 \<rho>\<^sub>1
begin
sublocale pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi>\<^sub>F H \<Phi>\<^sub>H map\<^sub>0 map\<^sub>1
apply unfold_locales
by (metis D.equivalence_maps_compose D.hcomp_in_hhomE \<rho>.components_are_equivalences
\<sigma>.components_are_equivalences map\<^sub>0_def map\<^sub>0_in_hom(1))
lemma is_pseudonatural_equivalence:
shows "pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
F \<Phi>\<^sub>F H \<Phi>\<^sub>H map\<^sub>0 map\<^sub>1"
..
end
locale pseudonatural_equivalence_whisker_right =
pseudonatural_transformation_whisker_right +
\<tau>: pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
begin
interpretation FoH: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H F \<Phi>\<^sub>F
..
interpretation GoH: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D H \<Phi>\<^sub>H G \<Phi>\<^sub>G
..
sublocale pseudonatural_equivalence V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>F o H\<close> FoH.cmp \<open>G o H\<close> GoH.cmp map\<^sub>0 map\<^sub>1
using map\<^sub>0_def \<tau>.components_are_equivalences
by unfold_locales simp
lemma is_pseudonatural_equivalence:
shows "pseudonatural_equivalence V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F o H) FoH.cmp (G o H) GoH.cmp map\<^sub>0 map\<^sub>1"
..
end
locale pseudonatural_equivalence_whisker_left =
pseudonatural_transformation_whisker_left +
\<tau>: pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
begin
interpretation HoF: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F H \<Phi>\<^sub>H
..
interpretation HoG: composite_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G H \<Phi>\<^sub>H
..
sublocale pseudonatural_equivalence V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>H o F\<close> HoF.cmp \<open>H o G\<close> HoG.cmp map\<^sub>0 map\<^sub>1
using map\<^sub>0_def \<tau>.components_are_equivalences H.preserves_equivalence_maps
by unfold_locales simp
lemma is_pseudonatural_equivalence:
shows "pseudonatural_equivalence V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(H o F) HoF.cmp (H o G) HoG.cmp map\<^sub>0 map\<^sub>1"
..
end
subsubsection "Converse of a Pseudonatural Equivalence"
text \<open>
It is easy to see that natural isomorphism between ordinary functors is a symmetric
relation because a unique inverse to a natural isomorphism is obtained merely by inverting
the components. However the situation is more difficult for pseudonatural equivalences
because they do not have unique inverses. Instead, we have to choose a quasi-inverse for
each of the components. In order to satisfy the required coherence conditions,
it is necessary for these quasi-inverses to be part of chosen adjoint equivalences.
Some long calculations to establish the coherence conditions seem unavoidable.
The purpose of this section is to carry out the construction, given a pseudonatural
equivalence, of a ``converse'' pseudonatural equivalence in the opposite direction.
\<close>
locale converse_pseudonatural_equivalence =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F +
G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G +
\<tau>: pseudonatural_equivalence
begin
abbreviation (input) F\<^sub>0
where "F\<^sub>0 \<equiv> F.map\<^sub>0"
abbreviation (input) G\<^sub>0
where "G\<^sub>0 \<equiv> G.map\<^sub>0"
definition map\<^sub>0
where "map\<^sub>0 a = (SOME g. \<exists>\<eta> \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) g \<eta> \<epsilon>)"
abbreviation (input) \<tau>\<^sub>0'
where "\<tau>\<^sub>0' \<equiv> map\<^sub>0"
definition unit
where "unit a = (SOME \<eta>. \<exists>\<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a) \<eta> \<epsilon>)"
abbreviation (input) \<eta>
where "\<eta> \<equiv> unit"
definition counit
where "counit a = (SOME \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a) (\<eta> a) \<epsilon>)"
abbreviation (input) \<epsilon>
where "\<epsilon> \<equiv> counit"
lemma chosen_adjoint_equivalence:
assumes "C.obj a"
shows "adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a) (\<eta> a) (\<epsilon> a)"
using assms \<tau>.components_are_equivalences map\<^sub>0_def unit_def counit_def
D.obj_is_equivalence_map D.equivalence_map_extends_to_adjoint_equivalence
someI_ex [of "\<lambda>g. \<exists>\<eta> \<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) g \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<eta>. \<exists>\<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a) \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<epsilon>. adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a) (\<eta> a) \<epsilon>"]
by simp
lemma map\<^sub>0_in_hhom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>\<tau>\<^sub>0' a : G\<^sub>0 a \<rightarrow>\<^sub>D F\<^sub>0 a\<guillemotright>"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis
using assms \<tau>.map\<^sub>0_in_hhom [of a] antipar by auto
qed
lemma map\<^sub>0_simps [simp]:
assumes "C.obj a"
shows "D.ide (\<tau>\<^sub>0' a)" and "src\<^sub>D (\<tau>\<^sub>0' a) = G\<^sub>0 a" and "trg\<^sub>D (\<tau>\<^sub>0' a) = F\<^sub>0 a"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "D.ide (\<tau>\<^sub>0' a)"
by simp
show "src\<^sub>D (\<tau>\<^sub>0' a) = G\<^sub>0 a"
using assms map\<^sub>0_in_hhom by blast
show "trg\<^sub>D (\<tau>\<^sub>0' a) = F\<^sub>0 a"
using assms map\<^sub>0_in_hhom by blast
qed
lemma equivalence_map_map\<^sub>0 [simp]:
assumes "C.obj a"
shows "D.equivalence_map (\<tau>\<^sub>0' a)"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "D.equivalence_map (\<tau>\<^sub>0' a)"
using D.equivalence_map_def adjoint_equivalence_in_bicategory_axioms
dual_equivalence
by blast
qed
lemma unit_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>\<eta> a : F\<^sub>0 a \<rightarrow>\<^sub>D F\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<eta> a : F\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a\<guillemotright>"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "\<guillemotleft>\<eta> a : F\<^sub>0 a \<rightarrow>\<^sub>D F\<^sub>0 a\<guillemotright>"
using assms unit_in_hom antipar(1) by auto
show "\<guillemotleft>\<eta> a : F\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a\<guillemotright>"
using assms unit_in_hom antipar(1) by auto
qed
lemma unit_simps [simp]:
assumes "C.obj a"
shows "D.iso (\<eta> a)" and "D.arr (\<eta> a)"
and "src\<^sub>D (\<eta> a) = F\<^sub>0 a" and "trg\<^sub>D (\<eta> a) = F\<^sub>0 a"
and "D.dom (\<eta> a) = F\<^sub>0 a" and "D.cod (\<eta> a) = \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "D.iso (\<eta> a)" and "D.arr (\<eta> a)"
and "src\<^sub>D (\<eta> a) = F\<^sub>0 a" and "trg\<^sub>D (\<eta> a) = F\<^sub>0 a"
and "D.dom (\<eta> a) = F\<^sub>0 a" and "D.cod (\<eta> a) = \<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a"
using assms by auto
qed
lemma iso_unit:
assumes "C.obj a"
shows "D.iso (\<eta> a)"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis by simp
qed
lemma counit_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>\<epsilon> a : G\<^sub>0 a \<rightarrow>\<^sub>D G\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<epsilon> a : \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a \<Rightarrow>\<^sub>D G\<^sub>0 a\<guillemotright>"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "\<guillemotleft>\<epsilon> a : G\<^sub>0 a \<rightarrow>\<^sub>D G\<^sub>0 a\<guillemotright>"
using assms counit_in_hom antipar(2) by auto
show " \<guillemotleft>\<epsilon> a : \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a \<Rightarrow>\<^sub>D G\<^sub>0 a\<guillemotright>"
using assms counit_in_hom antipar(2) by simp
qed
lemma counit_simps [simp]:
assumes "C.obj a"
shows "D.iso (\<epsilon> a)" and "D.arr (\<epsilon> a)"
and "src\<^sub>D (\<epsilon> a) = G\<^sub>0 a" and "trg\<^sub>D (\<epsilon> a) = G\<^sub>0 a"
and "D.dom (\<epsilon> a) = \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a" and "D.cod (\<epsilon> a) = G\<^sub>0 a"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "D.iso (\<epsilon> a)" and "D.arr (\<epsilon> a)"
and "src\<^sub>D (\<epsilon> a) = G\<^sub>0 a" and "trg\<^sub>D (\<epsilon> a) = G\<^sub>0 a"
and "D.dom (\<epsilon> a) = \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a" and "D.cod (\<epsilon> a) = G\<^sub>0 a"
using assms by auto
qed
lemma iso_counit:
assumes "C.obj a"
shows "D.iso (\<epsilon> a)"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis by simp
qed
lemma quasi_inverts_components:
assumes "C.obj a"
shows "D.isomorphic (\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) (F\<^sub>0 a)"
and "D.isomorphic (\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) (G\<^sub>0 a)"
and "D.quasi_inverses (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a)"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using assms chosen_adjoint_equivalence by simp
show "D.isomorphic (\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) (F\<^sub>0 a)"
using assms D.isomorphic_def D.isomorphic_symmetric unit_is_iso by blast
show "D.isomorphic (\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) (G\<^sub>0 a)"
using assms D.isomorphic_def counit_is_iso by blast
show "D.quasi_inverses (\<tau>\<^sub>0 a) (\<tau>\<^sub>0' a)"
using D.quasi_inverses_def equivalence_in_bicategory_axioms by auto
qed
definition map\<^sub>1
where "map\<^sub>1 f = (\<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> (src\<^sub>C f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 (src\<^sub>C f), \<tau>\<^sub>0' (src\<^sub>C f)]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' (trg\<^sub>C f), G f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f), \<tau>\<^sub>0' (src\<^sub>C f)] \<cdot>\<^sub>D
((\<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' (trg\<^sub>C f), \<tau>\<^sub>0 (trg\<^sub>C f), F f] \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f)) \<cdot>\<^sub>D
((\<eta> (trg\<^sub>C f) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f))"
abbreviation (input) \<tau>\<^sub>1'
where "\<tau>\<^sub>1' \<equiv> map\<^sub>1"
lemma map\<^sub>1_in_hom [intro]:
assumes "C.ide f"
shows "\<guillemotleft>\<tau>\<^sub>1' f : G\<^sub>0 (src\<^sub>C f) \<rightarrow>\<^sub>D F\<^sub>0 (trg\<^sub>C f)\<guillemotright>"
and "\<guillemotleft>\<tau>\<^sub>1' f : F f \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D G f\<guillemotright>"
proof -
show "\<guillemotleft>\<tau>\<^sub>1' f : F f \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D G f\<guillemotright>"
using assms \<tau>.iso_map\<^sub>1_ide \<tau>.map\<^sub>1_in_vhom
by (unfold map\<^sub>1_def, intro D.comp_in_homI) auto
thus "\<guillemotleft>\<tau>\<^sub>1' f : G\<^sub>0 (src\<^sub>C f) \<rightarrow>\<^sub>D F\<^sub>0 (trg\<^sub>C f)\<guillemotright>"
using assms D.vconn_implies_hpar(1-2) by auto
qed
lemma map\<^sub>1_simps [simp]:
assumes "C.ide f"
shows "D.arr (\<tau>\<^sub>1' f)"
and "src\<^sub>D (\<tau>\<^sub>1' f) = G\<^sub>0 (src\<^sub>C f)" and "trg\<^sub>D (\<tau>\<^sub>1' f) = F\<^sub>0 (trg\<^sub>C f)"
and "D.dom (\<tau>\<^sub>1' f) = F f \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f)" and "D.cod (\<tau>\<^sub>1' f) = \<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D G f"
using assms map\<^sub>1_in_hom by auto
lemma iso_map\<^sub>1_ide:
assumes "C.ide f"
shows "D.iso (\<tau>\<^sub>1' f)"
proof -
interpret src: adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>\<tau>\<^sub>0 (src\<^sub>C f)\<close> \<open>\<tau>\<^sub>0' (src\<^sub>C f)\<close> \<open>\<eta> (src\<^sub>C f)\<close> \<open>\<epsilon> (src\<^sub>C f)\<close>
using assms chosen_adjoint_equivalence by simp
interpret trg: adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>\<tau>\<^sub>0 (trg\<^sub>C f)\<close> \<open>\<tau>\<^sub>0' (trg\<^sub>C f)\<close> \<open>\<eta> (trg\<^sub>C f)\<close> \<open>\<epsilon> (trg\<^sub>C f)\<close>
using assms chosen_adjoint_equivalence by simp
show ?thesis
unfolding map\<^sub>1_def
using assms \<tau>.iso_map\<^sub>1_ide
by (intro D.isos_compose) auto
qed
interpretation EV: self_evaluation_map V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
notation EV.eval ("\<lbrace>_\<rbrace>")
sublocale pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G F \<Phi>\<^sub>F \<tau>\<^sub>0' \<tau>\<^sub>1'
proof
show "\<And>a. C.obj a \<Longrightarrow> D.ide (\<tau>\<^sub>0' a)"
using map\<^sub>0_simps(1) by simp
show "\<And>f. C.ide f \<Longrightarrow> D.iso (\<tau>\<^sub>1' f)"
using iso_map\<^sub>1_ide by simp
show "\<And>a. C.obj a \<Longrightarrow> \<guillemotleft>\<tau>\<^sub>0' a : src\<^sub>D (G a) \<rightarrow>\<^sub>D src\<^sub>D (F a)\<guillemotright>"
by fastforce
show "\<And>f. C.ide f \<Longrightarrow> \<guillemotleft>\<tau>\<^sub>1' f : F f \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f) \<Rightarrow>\<^sub>D \<tau>\<^sub>0' (trg\<^sub>C f) \<star>\<^sub>D G f\<guillemotright>"
by auto
show "\<And>a. C.obj a \<Longrightarrow> D.equivalence_map (\<tau>\<^sub>0' a)"
by simp
show "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> \<tau>\<^sub>1' (C.cod \<mu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C \<mu>))
= (\<tau>\<^sub>0' (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1' (C.dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "C.arr \<mu>"
let ?a = "src\<^sub>C \<mu>"
let ?b = "trg\<^sub>C \<mu>"
let ?f = "C.dom \<mu>"
let ?g = "C.cod \<mu>"
have "\<tau>\<^sub>1' (C.cod \<mu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?g \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?g] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?g) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?g] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F \<mu> \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
unfolding map\<^sub>1_def
using \<mu> D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?g \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?g] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<eta> ?b \<star>\<^sub>D F ?g) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((trg\<^sub>D (F \<mu>) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> D.whisker_right D.lunit'_naturality [of "F \<mu>"] D.comp_assoc
D.whisker_right [of "\<tau>\<^sub>0' ?a" "trg\<^sub>D (F \<mu>) \<star>\<^sub>D F \<mu>" "\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f]"]
by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?g \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?g] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "((\<eta> ?b \<star>\<^sub>D F ?g) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((trg\<^sub>D (F \<mu>) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "((\<eta> ?b \<star>\<^sub>D F ?g) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((trg\<^sub>D (F \<mu>) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<eta> ?b \<star>\<^sub>D F ?g) \<cdot>\<^sub>D (trg\<^sub>D (F \<mu>) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> D.whisker_right D.obj_simps by simp
also have "... = (\<eta> ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> D.interchange [of "\<eta> ?b" "trg\<^sub>D (F \<mu>)" "F ?g" "F \<mu>"]
D.comp_arr_dom D.comp_cod_arr
by simp
also have "... = ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D (\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> D.interchange [of "\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<eta> ?b" "F \<mu>" "F ?f"]
D.comp_arr_dom D.comp_cod_arr
by simp
also have "... = (((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> D.whisker_right by simp
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?g \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?g] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?g] \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> D.whisker_right by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> D.assoc_naturality [of "\<tau>\<^sub>0' ?b" "\<tau>\<^sub>0 ?b" "F \<mu>"] by simp
also have "... = ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> D.whisker_right by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?g] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?g \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> \<tau>.iso_map\<^sub>1_ide D.whisker_right by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g) \<cdot>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> \<tau>.iso_map\<^sub>1_ide \<tau>.naturality
D.invert_opposite_sides_of_square
[of "\<tau>\<^sub>1 ?g" "G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a" "\<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>" "\<tau>\<^sub>1 ?f"]
by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using \<mu> \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = ((\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> \<tau>.iso_map\<^sub>1_ide D.whisker_right by simp
finally have "((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?g)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F \<mu>) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> D.assoc_naturality [of "\<tau>\<^sub>0' ?b" "G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a" "\<tau>\<^sub>0' ?a"] D.comp_assoc
by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D ((G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D[G ?f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using \<mu> D.assoc_naturality [of "G \<mu>" "\<tau>\<^sub>0 ?a" "\<tau>\<^sub>0' ?a"] by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
using \<mu> D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?g, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D G\<^sub>0 ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<tau>\<^sub>0' ?b \<star>\<^sub>D (G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using \<mu> D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<epsilon> ?a"
using \<mu> D.interchange [of "G ?g" "G \<mu>" "\<epsilon> ?a" "\<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
D.comp_arr_dom D.comp_cod_arr
by simp
also have "... = \<tau>\<^sub>0' ?b \<star>\<^sub>D (G \<mu> \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D (G ?f \<star>\<^sub>D \<epsilon> ?a)"
using \<mu> D.interchange [of "G \<mu>" "G ?f" "G\<^sub>0 ?a" "\<epsilon> ?a"]
D.comp_arr_dom D.comp_cod_arr
by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?f \<star>\<^sub>D \<epsilon> ?a)"
using \<mu> D.whisker_left D.obj_simps by auto
finally have "(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?g \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?f \<star>\<^sub>D \<epsilon> ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G ?f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G ?f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G ?f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F ?f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D G\<^sub>0 ?a)
= \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g] \<cdot>\<^sub>D (G \<mu> \<star>\<^sub>D G\<^sub>0 ?a)"
using \<mu> D.whisker_left D.obj_simps by simp
also have "... = \<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<cdot>\<^sub>D \<r>\<^sub>D[G ?f]"
using \<mu> D.runit_naturality [of "G \<mu>"] by simp
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?f])"
using \<mu> D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?g]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu> \<star>\<^sub>D G\<^sub>0 ?a)
= (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G ?f])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1' ?f"
unfolding map\<^sub>1_def
using \<mu> D.comp_assoc by simp
finally show "\<tau>\<^sub>1' ?g \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D \<tau>\<^sub>0' ?a) = (\<tau>\<^sub>0' ?b \<star>\<^sub>D G \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1' ?f"
by blast
qed
show "\<And>a. C.obj a \<Longrightarrow>
(\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0' a] = \<tau>\<^sub>1' a \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
fix a
assume "C.obj a"
hence a: "C.obj a \<and> C.arr a \<and> C.ide a \<and> src\<^sub>C a = a \<and> trg\<^sub>C a = a"
by auto
have 1: "D.ide (F\<^sub>0 a)"
using a F.map\<^sub>0_simps(1) by blast
have 2: "D.ide (G\<^sub>0 a)"
using a G.map\<^sub>0_simps(1) by blast
have "\<tau>\<^sub>1' a \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a)"
unfolding map\<^sub>1_def
using a D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(((\<eta> a \<star>\<^sub>D F a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((F\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a) = \<l>\<^sub>D\<^sup>-\<^sup>1[F a] \<cdot>\<^sub>D F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.whisker_right by simp
also have "... = (F\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.lunit'_naturality [of "F.unit a"] by simp
also have "... = ((F\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a 1 D.whisker_right by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a) =
((F\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "((\<eta> a \<star>\<^sub>D F a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((F\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<eta> a \<star>\<^sub>D F a) \<cdot>\<^sub>D (F\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a 1 D.whisker_right by simp
also have "... = (\<eta> a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.interchange [of "\<eta> a" "F\<^sub>0 a" "F a" "F.unit a"]
D.comp_cod_arr D.comp_arr_dom
by simp
also have "... = ((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.interchange [of "\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a" "\<eta> a" "F.unit a" "F\<^sub>0 a"]
D.comp_cod_arr D.comp_arr_dom
by simp
also have "... = (((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a 1 D.whisker_right by simp
finally have "((\<eta> a \<star>\<^sub>D F a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((F\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= (((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F a] \<cdot>\<^sub>D ((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.whisker_right by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.assoc_naturality [of "\<tau>\<^sub>0' a" "\<tau>\<^sub>0 a" "F.unit a"] by simp
also have "... = ((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a 1 D.whisker_right by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= ((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, G a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "((\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a \<tau>.iso_map\<^sub>1_ide D.whisker_right by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a)) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a"
proof -
have "(\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a] = \<tau>\<^sub>1 a \<cdot>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a)"
using a \<tau>.respects_unit by simp
hence "D.inv (\<tau>\<^sub>1 a) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0 a]
= G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a"
using a \<tau>.iso_map\<^sub>1_ide D.invert_side_of_triangle(1) by simp
hence "D.inv (\<tau>\<^sub>1 a) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]
= (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]"
using a D.iso_lunit D.invert_side_of_triangle(2) D.comp_assoc by simp
hence "D.inv (\<tau>\<^sub>1 a) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a)
= (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]"
using a D.iso_runit D.comp_assoc
D.invert_side_of_triangle(2)
[of "(G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]" "D.inv (\<tau>\<^sub>1 a) \<cdot>\<^sub>D (\<tau>\<^sub>0 a \<star>\<^sub>D F.unit a)"
"\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]"]
by simp
thus ?thesis by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.whisker_left by simp
also have "... = ((\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a D.whisker_right by simp
finally have "((\<tau>\<^sub>0' a \<star>\<^sub>D D.inv (\<tau>\<^sub>1 a)) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D F.unit a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= ((\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a D.assoc_naturality [of "\<tau>\<^sub>0' a" "G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a" "\<tau>\<^sub>0' a"] D.comp_assoc
by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D ((G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a D.whisker_left by simp
also have "... = \<tau>\<^sub>0' a \<star>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]"
using a D.assoc_naturality [of "G.unit a" "\<tau>\<^sub>0 a" "\<tau>\<^sub>0' a"] by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a])"
using a 2 D.whisker_left [of "\<tau>\<^sub>0' a"] by simp
finally have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D (G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D G\<^sub>0 a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<epsilon> a"
using a D.whisker_left D.comp_arr_dom D.comp_cod_arr
D.interchange [of "G a" "G.unit a" "\<epsilon> a" "\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a"]
by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D G\<^sub>0 a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a)"
using a 2 D.whisker_left D.comp_arr_dom D.comp_cod_arr
D.interchange [of "G.unit a" "G\<^sub>0 a" "G\<^sub>0 a" "\<epsilon> a"]
by simp
finally have "(\<tau>\<^sub>0' a \<star>\<^sub>D G a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D G\<^sub>0 a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, G\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D G\<^sub>0 a)
= \<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a] \<cdot>\<^sub>D (G.unit a \<star>\<^sub>D G\<^sub>0 a)"
using a 2 D.whisker_left by simp
also have "... = \<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<cdot>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]"
using a D.runit_naturality [of "G.unit a"] by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a])"
using a 2 D.whisker_left by simp
finally have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a \<star>\<^sub>D G\<^sub>0 a)
= (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' a, G\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a 1 D.whisker_right D.runit_hcomp(1) D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a D.assoc_naturality [of "\<tau>\<^sub>0' a" "\<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a]" "\<tau>\<^sub>0' a"] D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a])) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a 2 D.whisker_left D.lunit_hcomp(4) D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
((\<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a])
= \<tau>\<^sub>0' a \<star>\<^sub>D (G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a]"
using a 2 D.whisker_left by simp
also have "... = \<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a] \<cdot>\<^sub>D \<epsilon> a"
using a D.lunit'_naturality [of "\<epsilon> a"] by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<epsilon> a)"
using a 2 D.whisker_left by simp
finally have "(\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a \<star>\<^sub>D \<tau>\<^sub>0' a])
= (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<epsilon> a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<eta> a \<star>\<^sub>D \<tau>\<^sub>0' a)) \<cdot>\<^sub>D
(\<r>\<^sub>D[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
have "(\<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a"
using a 1 D.whisker_right by simp
also have "... = \<eta> a \<cdot>\<^sub>D \<r>\<^sub>D[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a"
using a D.runit_naturality [of "\<eta> a"] by simp
also have "... = (\<eta> a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (\<r>\<^sub>D[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
using a 1 D.whisker_right by simp
finally have "(\<r>\<^sub>D[\<tau>\<^sub>0' a \<star>\<^sub>D \<tau>\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D ((\<eta> a \<star>\<^sub>D F\<^sub>0 a) \<star>\<^sub>D \<tau>\<^sub>0' a)
= (\<eta> a \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (\<r>\<^sub>D[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]) \<cdot>\<^sub>D
\<l>\<^sub>D[\<tau>\<^sub>0' a] \<cdot>\<^sub>D
(\<r>\<^sub>D[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a)"
proof -
interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>\<tau>\<^sub>0 a\<close> \<open>\<tau>\<^sub>0' a\<close> \<open>\<eta> a\<close> \<open>\<epsilon> a\<close>
using a chosen_adjoint_equivalence by simp
have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<epsilon> a) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' a, \<tau>\<^sub>0 a, \<tau>\<^sub>0' a] \<cdot>\<^sub>D (\<eta> a \<star>\<^sub>D \<tau>\<^sub>0' a)
= \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0' a]"
using triangle_right by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0' a]"
proof -
have "\<l>\<^sub>D[\<tau>\<^sub>0' a] \<cdot>\<^sub>D (\<r>\<^sub>D[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \<star>\<^sub>D \<tau>\<^sub>0' a) = \<l>\<^sub>D[\<tau>\<^sub>0' a]"
proof -
have "D.seq \<r>\<^sub>D[F\<^sub>0 a] \<l>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a]"
using a 1 by auto
thus ?thesis
using a D.unitor_coincidence [of "F\<^sub>0 a"] D.comp_arr_dom D.whisker_right
by (metis D.comp_arr_inv' D.iso_runit D.lunit_simps(1,4)
D.objE D.runit_simps(5) F.map\<^sub>0_simps(1) map\<^sub>0_simps(1,3))
qed
moreover have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]
= \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]"
proof -
have "(\<tau>\<^sub>0' a \<star>\<^sub>D \<r>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a] =
((\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D[G\<^sub>0 a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a])) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]"
using a D.unitor_coincidence D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D \<l>\<^sub>D[G\<^sub>0 a] \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]"
using a 2 D.whisker_left by simp
also have "... = (\<tau>\<^sub>0' a \<star>\<^sub>D G\<^sub>0 a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]"
using a 2 D.comp_arr_inv' by simp
also have "... = \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a]"
using a 2 D.comp_cod_arr by simp
finally show ?thesis by blast
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
finally show "(\<tau>\<^sub>0' a \<star>\<^sub>D G.unit a) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' a] \<cdot>\<^sub>D \<l>\<^sub>D[\<tau>\<^sub>0' a]
= \<tau>\<^sub>1' a \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D \<tau>\<^sub>0' a)"
by simp
qed
show "\<And>f g. \<lbrakk>C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\<rbrakk> \<Longrightarrow>
(\<tau>\<^sub>0' (trg\<^sub>C g) \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' (trg\<^sub>C g), G g, G f] \<cdot>\<^sub>D
(\<tau>\<^sub>1' g \<star>\<^sub>D G f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0' (src\<^sub>C g), G f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1' f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' (src\<^sub>C f)]
= \<tau>\<^sub>1' (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' (src\<^sub>C f))"
proof -
fix f g
assume f: "C.ide f" and g: "C.ide g" and fg: "src\<^sub>C g = trg\<^sub>C f"
let ?a = "src\<^sub>C f"
let ?b = "trg\<^sub>C f"
let ?c = "trg\<^sub>C g"
(*
* TODO: The following are extremely problematic.
* I don't think all the cases are used, but they don't get discovered by auto
* and sledgehammer doesn't find them, either. Note that the same issue occurred
* for the previous subgoal, where I needed D.ide (F\<^sub>0 a) and D.ide (G\<^sub>0 a).
*)
have *: "D.ide (G\<^sub>0 ?a) \<and> D.ide (G\<^sub>0 ?b) \<and> D.ide (G\<^sub>0 ?b) \<and> D.ide (G\<^sub>0 ?c)"
using f g C.ideD(1) G.map\<^sub>0_simps(1) by blast
have **: "D.ide (F\<^sub>0 ?a) \<and> D.ide (F\<^sub>0 ?b) \<and> D.ide (F\<^sub>0 ?b) \<and> D.ide (F\<^sub>0 ?c)"
using f g C.ideD(1) F.map\<^sub>0_simps(1) by blast
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
(\<tau>\<^sub>1' g \<star>\<^sub>D G f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0' (src\<^sub>C g), G f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>1' f) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b)
\<star>\<^sub>D G f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
unfolding map\<^sub>1_def
using f g fg D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, \<tau>\<^sub>0' ?b, G f]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b)
\<star>\<^sub>D G f
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f)"
proof -
have 1: "D.arr ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b))"
using f g fg \<tau>.iso_map\<^sub>1_ide by auto
thus ?thesis
using f g fg D.whisker_right [of "G f"]
by (metis D.seqE G.preserves_ide)
qed
moreover have "F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have 1: "D.arr ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a))"
using f g fg \<tau>.iso_map\<^sub>1_ide map\<^sub>1_def map\<^sub>1_simps(1) by presburger
thus ?thesis
using f g fg D.whisker_left [of "F g"]
by (metis D.seqE F.preserves_ide)
qed
ultimately show ?thesis
using f g fg D.comp_assoc by simp
qed
(*
* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] and \<l>\<^sub>D\<^sup>-\<^sup>1[F f] to the bottom, followed by \<eta> ?c and D.inv (\<tau>\<^sub>1 g).
* Move \<epsilon> ?b down across D.inv (\<tau>\<^sub>1 g) to meet \<eta> ?b, where they can be canceled.
*)
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, \<tau>\<^sub>0' ?b, G f]
= \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f)"
using f g fg D.assoc'_naturality [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "\<tau>\<^sub>0' ?b" "G f"] by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f" "\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "(\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(F g \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D\<^sup>-\<^sup>1[F g] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f)) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f]
= \<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f)"
using f g fg D.assoc'_naturality [of "\<eta> ?c \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b" "G f"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f) \<cdot>\<^sub>D ((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f" "\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?c down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f \<star>\<^sub>D G\<^sub>0 ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f \<star>\<^sub>D G\<^sub>0 ?a]"
using f g fg D.assoc'_naturality [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b" "\<r>\<^sub>D[G f]"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f \<star>\<^sub>D G\<^sub>0 ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg D.assoc'_naturality [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg
D.assoc'_naturality [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide
D.assoc_naturality [of "\<tau>\<^sub>0' ?b" "D.inv (\<tau>\<^sub>1 f)" "\<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.assoc'_naturality
[of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<tau>\<^sub>0' ?b" "D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f" "\<r>\<^sub>D[G f]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<r>\<^sub>D[G f]" "G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D G\<^sub>0 ?a" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<epsilon> ?a" "G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
"(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f" "\<r>\<^sub>D[G f]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<r>\<^sub>D[G f]" "G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D G\<^sub>0 ?a" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<epsilon> ?a" "G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]" "(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.assoc'_naturality [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]" "\<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g" "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g" "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide
D.assoc'_naturality
[of "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)" "\<tau>\<^sub>0' ?b" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)" "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g"
"\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<tau>\<^sub>0' ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 g) down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)" "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g"
"(\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?b up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<cdot>\<^sub>D (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D
(((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]" "\<eta> ?c \<star>\<^sub>D F g"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<cdot>\<^sub>D (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D ((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_cod_arr by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<cdot>\<^sub>D (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.interchange by auto
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?b up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D ((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_cod_arr
D.interchange
[of "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)" "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]"
"((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom
D.interchange
[of "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)" "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]"
"(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D ((\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move \<epsilon> ?b down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f"
using f g fg D.whisker_right by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f"
using f g fg D.pentagon
D.invert_side_of_triangle(2)
[of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"]
D.comp_assoc
by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f)"
using f g fg D.whisker_right by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<epsilon> ?b down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f"
using f g fg D.whisker_right [of "G f"] by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f"
using f g fg D.assoc_naturality [of "\<tau>\<^sub>0' ?c" "G g" "\<epsilon> ?b"] by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)"
using f g fg * D.whisker_right [of "G f"] by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move \<epsilon> ?b down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
using f g fg D.pentagon by simp
hence "\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b)"
using f g fg D.comp_assoc D.invert_side_of_triangle(1) by simp
hence "\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b)
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
using f g fg D.comp_assoc
D.invert_side_of_triangle(2)
[of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"]
by simp
hence "\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f"
using f g fg D.whisker_right by simp
hence "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f)
= (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)"
using f g fg D.whisker_right by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<epsilon> ?b down *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f"
using f g fg D.whisker_right by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f"
using f g fg D.assoc'_naturality [of "\<tau>\<^sub>0' ?c" "G g" "\<epsilon> ?b"] by simp
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)"
using f g fg * D.whisker_right by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f)
= (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* There is a cancellation of associativities here. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f"
using f g fg * D.comp_arr_inv' D.whisker_right
by (metis C.ideD(1) C.obj_trg D.comp_assoc_assoc'(1) D.hcomp_simps(2) D.hseqI'
D.ideD(1) G.map\<^sub>0_simps(3) G.preserves_ide G.preserves_src G.preserves_trg
map\<^sub>0_simps(1-2))
moreover have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f"
using f g fg D.comp_cod_arr by simp
ultimately have "((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<r>\<^sub>D[G f] up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f" "\<r>\<^sub>D[G f]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<r>\<^sub>D[G f]" "G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using f g fg D.comp_assoc by simp
qed
(* Move \<r>\<^sub>D[G f] up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f" "\<r>\<^sub>D[G f]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<r>\<^sub>D[G f]" "G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<r>\<^sub>D[G f] up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f" "\<r>\<^sub>D[G f]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b"
"\<r>\<^sub>D[G f]" "G f \<star>\<^sub>D G\<^sub>0 ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<epsilon> ?a up. The useful effect is on the associativity part of the term. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D G\<^sub>0 ?a" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"G f \<star>\<^sub>D \<epsilon> ?a" "G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<epsilon> ?a up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D G\<^sub>0 ?a" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"G f \<star>\<^sub>D \<epsilon> ?a" "G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<epsilon> ?a up and \<epsilon> ?b down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b"
"\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D G\<^sub>0 ?a" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b"
"\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b"
"G f \<star>\<^sub>D \<epsilon> ?a" "G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move \<epsilon> ?b down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
using f g fg D.pentagon by simp
hence "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b)"
using f g fg D.comp_assoc
D.invert_side_of_triangle(2)
[of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"]
by simp
hence "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]
\<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b)
\<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"
by simp
thus ?thesis
using f g fg D.whisker_right by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<epsilon> ?b down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b]
\<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.whisker_right by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b)
\<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.assoc_naturality [of "\<tau>\<^sub>0' ?c" "G g" "\<epsilon> ?b"] by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg * D.whisker_right by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 f) up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]" "(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 f) up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]" "(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Prepare to move D.inv (\<tau>\<^sub>1 f) up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]" "(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 f) up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 f) up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b]"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move D.inv (\<tau>\<^sub>1 f) up across \<epsilon> ?b. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b"
"(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<eta> ?b up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)"
"(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(*
* Introduce associativities to align (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) with
* (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) in the middle.
*)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom by simp
also have "... = ((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]\<rbrace>"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... =
\<lbrace>((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg
D.assoc'_naturality [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g" "\<epsilon> ?b" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
using f g fg * D.comp_arr_dom D.comp_cod_arr D.whisker_left
D.interchange [of "\<epsilon> ?b" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b)" "(\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg * D.comp_arr_dom D.comp_cod_arr D.whisker_left
D.interchange [of "G\<^sub>0 ?b" "\<epsilon> ?b" "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"
"\<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D
(\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg D.whisker_left by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D
\<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.assoc_naturality [of "\<epsilon> ?b" "\<tau>\<^sub>0 ?b" "F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg * D.whisker_left by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis by simp
qed
moreover
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.comp_arr_dom by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= \<lbrace>(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>])\<rbrace>"
- using f g fg ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = \<lbrace>(\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D
((\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"
using f g fg ** D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b]"
"(\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.assoc'_naturality [of "\<eta> ?b" "F f" "\<tau>\<^sub>0' ?a"]
by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b]" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b"
"\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"
"\<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D (\<eta> ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.assoc'_naturality [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g" "\<tau>\<^sub>0 ?b"
"\<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by force
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D
(\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg ** D.whisker_left by auto
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.assoc_naturality [of "\<tau>\<^sub>0 ?b" "\<eta> ?b" "F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by auto
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.whisker_left by auto
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis by blast
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
(* Cancel out all the intervening associativities. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b] \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<tau>\<^sub>0' ?b) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (\<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= \<lbrace>((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>])\<rbrace>"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... =
\<lbrace>(\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?b\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally show ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
(* Hooray! We can finally cancel \<eta> ?b with \<epsilon> ?b using the triangle identity. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 ?b] \<star>\<^sub>D (F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D
((\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.whisker_left by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D
(\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<cdot>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b)
\<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.whisker_right by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 ?b] \<star>\<^sub>D (F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
interpret adjoint_equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<tau>\<^sub>0 ?b\<close> \<open>\<tau>\<^sub>0' ?b\<close> \<open>\<eta> ?b\<close> \<open>\<epsilon> ?b\<close>
using f chosen_adjoint_equivalence by simp
show ?thesis
using fg triangle_left by simp
qed
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<epsilon> ?b \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, \<tau>\<^sub>0' ?b, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D \<eta> ?b) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 ?b] \<star>\<^sub>D (F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Simplify some more canonical arrows. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 ?b] \<star>\<^sub>D (F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= (\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, G\<^sub>0 ?b, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G\<^sub>0 ?b, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b] \<cdot>\<^sub>D \<r>\<^sub>D[\<tau>\<^sub>0 ?b] \<star>\<^sub>D (F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F\<^sub>0 ?b \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>G\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>G\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>G\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<r>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> (\<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>)) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>])\<rbrace>"
- using f g fg * ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char D.\<ll>_ide_simp D.\<rr>_ide_simp
+ using f g fg * ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.\<ll>_ide_simp D.\<rr>_ide_simp
by simp
also have "... = \<lbrace>(\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^bold>[\<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, (\<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> (\<^bold>\<langle>F\<^sub>0 ?b\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>)\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = (\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
- using f g fg * ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char D.\<ll>_ide_simp D.\<rr>_ide_simp
+ using f g fg * ** D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C D.\<ll>_ide_simp D.\<rr>_ide_simp
by simp
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D G f) \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<r>\<^sub>D[G f])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]" "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b" "G f" "\<r>\<^sub>D[G f]"]
by simp
moreover have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg * D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b]" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b"
"G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
(*
* Move \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] and \<l>\<^sub>D[F f] outside, to get rid of
* G\<^sub>0 ?b and F\<^sub>0 ?b.
*)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.assoc_naturality [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g" "\<tau>\<^sub>0 ?b" "\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b]"
"\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b]" "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b"
"F f \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)"
"\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)" "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g"
"F f \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<tau>\<^sub>0 ?b) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b" "\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g]"
"D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = (\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g]" "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g" "(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
"(D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]"]
D.whisker_left
by simp (* 12 sec *)
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= (\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Introduce associativities to achieve (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) and (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)). *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom by simp
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left D.comp_assoc by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_inv' by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_cod_arr by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_inv_arr' by simp
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide
D.assoc'_naturality [of "\<tau>\<^sub>0' ?c" "G g" "D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])
= \<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.assoc_naturality [of "G g" "D.inv (\<tau>\<^sub>1 f)" "\<tau>\<^sub>0' ?a"]
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis by simp
qed
moreover have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom by auto
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg D.whisker_left D.comp_assoc by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (\<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg D.comp_arr_inv' by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg D.comp_cod_arr by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_inv_arr' by simp
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g)) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide
D.assoc'_naturality [of "\<tau>\<^sub>0' ?c" "D.inv (\<tau>\<^sub>1 g)" "F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a])
= \<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.assoc_naturality [of "D.inv (\<tau>\<^sub>1 g)" "F f" "\<tau>\<^sub>0' ?a"]
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis
by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
(* Cancel intervening associativities. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= \<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, (\<tau>\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g, \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, \<tau>\<^sub>0 ?b] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D \<tau>\<^sub>0 ?b, F f, \<tau>\<^sub>0' ?a])
= \<lbrace>(\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, (\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>G g\<^bold>\<rangle>,\<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>G g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>G g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>])\<rbrace>"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = \<lbrace>\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>G g\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?b\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
(* Apply "respects composition". *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<tau>\<^sub>0' ?c \<star>\<^sub>D ((G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D
(G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f)
\<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_right by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D
\<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<cdot>\<^sub>D
(D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f])
\<star>\<^sub>D \<tau>\<^sub>0' ?a"
proof -
have 1: "(\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f)
= D.inv ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f]) \<cdot>\<^sub>D
\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a]"
proof -
have "((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a]
= \<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a)"
using f g fg \<tau>.respects_hcomp D.comp_assoc by simp
hence "((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f]) \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f)
= \<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a]"
using f g fg D.comp_assoc
D.invert_side_of_triangle(2)
[of "\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a)"
"(\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<cdot>\<^sub>D
(\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f)"
"\<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a]"]
by fastforce
moreover have "D.iso ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f])"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_simp F.cmp_components_are_iso
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp F.cmp_components_are_iso
F.FF_def
by (intro D.isos_compose D.seqI D.hseqI') auto
moreover have "D.seq (\<tau>\<^sub>1 (g \<star>\<^sub>C f))
((\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a])"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp G.FF_def
by (intro D.seqI) auto
ultimately show ?thesis
using f g fg \<tau>.iso_map\<^sub>1_ide
D.invert_side_of_triangle(1)
[of "\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a]"
"(\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f]"
"(\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f)"]
by simp
qed
have "(G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f)
= D.inv ((\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f))"
proof -
have "D.inv ((\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f))
= D.inv (\<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f)) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 g \<star>\<^sub>D F f)"
proof -
have "D.iso ((\<tau>\<^sub>1 g \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f))"
using f g fg \<tau>.iso_map\<^sub>1_ide
by (intro D.isos_compose) auto
moreover have "D.iso (\<tau>\<^sub>1 g \<star>\<^sub>D F f)"
using f g fg \<tau>.iso_map\<^sub>1_ide by auto
ultimately show ?thesis
using \<tau>.iso_map\<^sub>1_ide D.inv_comp_left by simp
qed
also have "... = ((G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f]) \<cdot>\<^sub>D
(D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f)"
proof -
have "D.iso (\<a>\<^sub>D\<^sup>-\<^sup>1[G g, \<tau>\<^sub>0 (src\<^sub>C g), F f] \<cdot>\<^sub>D (G g \<star>\<^sub>D \<tau>\<^sub>1 f))"
using f g fg \<tau>.iso_map\<^sub>1_ide
by (intro D.isos_compose) auto
thus ?thesis
using f g fg \<tau>.iso_map\<^sub>1_ide D.inv_comp_left by simp
qed
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = D.inv (D.inv ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f]) \<cdot>\<^sub>D
\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a])"
using 1 by simp
also have "... = D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f])"
proof -
have 2: "D.iso ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f])"
- using f g fg \<tau>.iso_map\<^sub>1_ide C.VV.arr_char C.VV.dom_simp F.cmp_components_are_iso
+ using f g fg \<tau>.iso_map\<^sub>1_ide C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp F.cmp_components_are_iso
F.FF_def
by (intro D.isos_compose) auto
moreover have "D.iso (D.inv ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f]) \<cdot>\<^sub>D
\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a])"
- using 2 f g fg \<tau>.iso_map\<^sub>1_ide C.VV.arr_char C.VV.dom_simp
+ using 2 f g fg \<tau>.iso_map\<^sub>1_ide C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp
C.VV.cod_simp F.FF_def G.FF_def D.inv_comp_left
by (intro D.isos_compose) auto (* 10 sec *)
ultimately show ?thesis
using D.inv_comp_left
[of "D.inv ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f])"
"\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a]"]
D.inv_inv D.iso_inv_iso
by metis
qed
also have "... = \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f])"
proof -
have "D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a])
= \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))"
proof -
have 2: "D.iso (\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a])"
- using f g fg \<tau>.iso_map\<^sub>1_ide C.VV.arr_char C.VV.dom_simp
+ using f g fg \<tau>.iso_map\<^sub>1_ide C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp
C.VV.cod_simp G.FF_def
by (intro D.isos_compose D.seqI D.hseqI') auto
have "D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a])
= D.inv ((\<Phi>\<^sub>G (g, f) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[G g, G f, \<tau>\<^sub>0 ?a]) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))"
- using 2 f g fg * \<tau>.iso_map\<^sub>1_ide G.cmp_components_are_iso C.VV.arr_char
+ using 2 f g fg * \<tau>.iso_map\<^sub>1_ide G.cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
C.VV.dom_simp C.VV.cod_simp G.FF_def D.inv_comp_left
by simp
also have "... = (\<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a)) \<cdot>\<^sub>D
D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))"
- using 2 f g fg * \<tau>.iso_map\<^sub>1_ide G.cmp_components_are_iso C.VV.arr_char
+ using 2 f g fg * \<tau>.iso_map\<^sub>1_ide G.cmp_components_are_iso C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
C.VV.dom_simp C.VV.cod_simp G.FF_def D.inv_comp_left
by simp
also have "... = \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))"
using D.comp_assoc by simp
finally show ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
finally have "(G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<cdot>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f)
= \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<cdot>\<^sub>D (\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f]"
by blast
thus ?thesis by simp
qed
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D
(\<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_right
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.cmp_components_are_iso
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp G.cmp_components_are_iso
F.FF_def G.FF_def
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide D.whisker_left
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.cmp_components_are_iso
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp G.cmp_components_are_iso
F.FF_def G.FF_def
by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, \<tau>\<^sub>0 ?b, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<tau>\<^sub>1 g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D[F f] down, where it will cancel. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g" "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]"
"\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
D.whisker_right
by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]" "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g"
"F f \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
D.whisker_right
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<l>\<^sub>D[F f] down. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g" "\<eta> ?c \<star>\<^sub>D F g"
"\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a" "(F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = ((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_arr_dom D.comp_cod_arr
D.interchange [of "\<eta> ?c \<star>\<^sub>D F g" "F\<^sub>0 ?c \<star>\<^sub>D F g" "F f \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D (F\<^sub>0 ?b \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Cancel \<l>\<^sub>D[F f] and \<l>\<^sub>D\<^sup>-\<^sup>1[F f]. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) =
\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D (\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.comp_cod_arr
D.interchange [of "F\<^sub>0 ?c \<star>\<^sub>D F g" "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "\<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a" "\<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D[F f] \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.whisker_right by simp
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f D.comp_arr_inv' by simp
finally have "((F\<^sub>0 ?c \<star>\<^sub>D F g) \<star>\<^sub>D \<l>\<^sub>D[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<r>\<^sub>D\<^sup>-\<^sup>1[G g] up, where it will cancel. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg * D.comp_arr_dom
D.interchange [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b]" "\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g]"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]" "(G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.runit_hcomp(4) [of "\<tau>\<^sub>0' ?c" "G g"] by simp
finally have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G\<^sub>0 ?b] \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<r>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D G g] \<star>\<^sub>D (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move \<r>\<^sub>D\<^sup>-\<^sup>1[G g] up. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D (G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.comp_cod_arr
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b" "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]"
"G f \<star>\<^sub>D \<epsilon> ?a" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
using f g fg D.comp_arr_dom
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]" "\<tau>\<^sub>0' ?c \<star>\<^sub>D G g" "G f \<star>\<^sub>D \<epsilon> ?a"
"\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G\<^sub>0 ?b) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Cancel \<r>\<^sub>D\<^sup>-\<^sup>1[G g] and \<r>\<^sub>D[G g]. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f] \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> ?a)"
using f g fg D.interchange by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f] \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> ?a)"
using f g fg D.whisker_left by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f] \<cdot>\<^sub>D (G f \<star>\<^sub>D \<epsilon> ?a)"
using g D.comp_arr_inv' by simp
also have "... = ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)"
using f g fg D.whisker_left by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g]) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[G g]) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D ((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, (G f \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
using f g fg
D.assoc'_naturality [of "\<tau>\<^sub>0' ?c" "G g" "\<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg D.pentagon by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
using f g fg D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<a>\<^sub>D[G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D G\<^sub>0 ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D G\<^sub>0 ?a] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a)"
using f g fg D.assoc'_naturality [of "\<tau>\<^sub>0' ?c" "G g" "G f \<star>\<^sub>D \<epsilon> ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Move associativity to where it can be canceled. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<r>\<^sub>D[G f])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g) \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f \<star>\<^sub>D G\<^sub>0 ?a]
= \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<r>\<^sub>D[G f])"
using f g fg D.assoc'_naturality [of "\<tau>\<^sub>0' ?c" "G g" "\<r>\<^sub>D[G f]"] by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Cancel associativities. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c, G g, G f]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<r>\<^sub>D[G f])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<r>\<^sub>D[G f])"
using f g fg D.comp_arr_inv' D.comp_cod_arr by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Permute associativity to continue forming G g \<star>\<^sub>D G f. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D \<r>\<^sub>D[G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, G\<^sub>0 ?a])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= \<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a]"
using f g fg D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, G\<^sub>0 ?a] \<cdot>\<^sub>D ((G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a)"
using f g fg D.assoc_naturality [of "G g" "G f" "\<epsilon> ?a"] by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, G\<^sub>0 ?a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a)"
using f g fg * D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D G g \<star>\<^sub>D G f \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a])
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g, G f, G\<^sub>0 ?a]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Form \<r>\<^sub>D[G g \<star>\<^sub>D G f]. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g \<star>\<^sub>D G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
using f g fg * D.runit_hcomp D.whisker_left D.comp_assoc by simp
(* Move D.inv (\<Phi>\<^sub>G (g, f)) to the top and cancel with its inverse. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g \<star>\<^sub>D G f]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<tau>\<^sub>0' ?c \<star>\<^sub>D
\<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg G.cmp_components_are_iso D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D
(D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]"
using f g fg G.cmp_components_are_iso
D.assoc_naturality [of "D.inv (\<Phi>\<^sub>G (g, f))" "\<tau>\<^sub>0 ?a" "\<tau>\<^sub>0' ?a"]
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
using f g fg G.cmp_components_are_iso D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G g \<star>\<^sub>D G f, \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g \<star>\<^sub>D G f]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<tau>\<^sub>0' ?c \<star>\<^sub>D ((G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg G.cmp_components_are_iso D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<epsilon> ?a)"
using f g fg * D.comp_arr_dom D.comp_cod_arr
D.interchange [of "G g \<star>\<^sub>D G f" "D.inv (\<Phi>\<^sub>G (g, f))" "\<epsilon> ?a" "\<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a"]
by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D (G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a)"
using f g fg * D.comp_arr_dom D.comp_cod_arr
D.interchange [of "D.inv (\<Phi>\<^sub>G (g, f))" "G (g \<star>\<^sub>C f)" "G\<^sub>0 ?a" "\<epsilon> ?a"]
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a)"
using f g fg * G.cmp_components_are_iso D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D (G g \<star>\<^sub>D G f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D \<tau>\<^sub>0 ?a \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)))) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)])) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g \<star>\<^sub>D G f]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a)
= \<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g \<star>\<^sub>D G f] \<cdot>\<^sub>D (D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a)"
using f g fg * G.cmp_components_are_iso D.whisker_left by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]"
using f g fg * G.cmp_components_are_iso D.runit_naturality [of "D.inv (\<Phi>\<^sub>G (g, f))"]
by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)])"
using f g fg * G.cmp_components_are_iso D.whisker_left by simp
finally have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G g \<star>\<^sub>D G f]) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)) \<star>\<^sub>D G\<^sub>0 ?a) =
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)])"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)))) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)])
= \<tau>\<^sub>0' ?c \<star>\<^sub>D (\<Phi>\<^sub>G (g, f) \<cdot>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f))) \<cdot>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]"
using f g fg D.whisker_left G.cmp_components_are_iso
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
by simp
also have "... = \<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]"
using f g fg G.cmp_components_are_iso D.comp_arr_inv' D.comp_cod_arr by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D (\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<Phi>\<^sub>G (g, f)))) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)])
= \<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Perform similar manipulations on the bottom part. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[F\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.assoc_naturality [of "\<l>\<^sub>D\<^sup>-\<^sup>1[F g]" "F f" "\<tau>\<^sub>0' ?a"] D.comp_assoc by simp
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D \<a>\<^sub>D[F\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg **
D.assoc_naturality [of "\<eta> ?c \<star>\<^sub>D F g" "F f" "\<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg D.assoc_naturality [of "\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g]" "F f" "\<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
(* Replace \<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f by \<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)]. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= \<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.lunit_hcomp(2) [of "F g" "F f"] by simp
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.whisker_right [of "\<tau>\<^sub>0' ?a"] by simp
finally have "(\<l>\<^sub>D\<^sup>-\<^sup>1[F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a
= (\<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Move the associativity up to the big block of associativities. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg ** D.whisker_right by simp
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c, F g, F f] \<cdot>\<^sub>D (\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg ** D.assoc'_naturality [of "\<eta> ?c" "F g" "F f"] by simp
also have "... = (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg ** D.whisker_right by simp
finally have "(((\<eta> ?c \<star>\<^sub>D F g) \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* Simplify the block of associativities. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f \<star>\<^sub>D \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D F g, F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g] \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c, F g, F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= \<lbrace>(\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle>, \<^bold>\<langle>F g\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F g\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F g\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F g\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle>, \<^bold>\<langle>F g\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle>, \<^bold>\<langle>F g\<^bold>\<rangle>, \<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>)\<rbrace>"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>\<tau>\<^sub>0' ?c\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<^sub>0 ?c\<^bold>\<rangle>, \<^bold>\<langle>F g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>F f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>\<tau>\<^sub>0' ?a\<^bold>\<rangle>)\<rbrace>"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
- using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char D.VVV.arr_char
- D.VV.ide_char D.VV.arr_char
+ using f g fg D.\<alpha>_def D.\<alpha>'.map_ide_simp D.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ D.VV.ide_char\<^sub>S\<^sub>b\<^sub>C D.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
(* Permute \<Phi>\<^sub>F (g, f) to the bottom. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F (g \<star>\<^sub>C f), \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D ((\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F g \<star>\<^sub>D F f, \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F (g \<star>\<^sub>C f), \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.assoc_naturality [of "\<tau>\<^sub>0' ?c" "\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)" "\<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F (g \<star>\<^sub>C f), \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D (\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f))) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.whisker_right
by simp
also have "... = \<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.assoc_naturality [of "\<tau>\<^sub>0' ?c" "\<tau>\<^sub>0 ?c" "\<Phi>\<^sub>F (g, f)"]
by simp
also have "... = (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.whisker_right
by simp
finally have "((\<tau>\<^sub>0' ?c \<star>\<^sub>D (\<tau>\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f))) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F (g \<star>\<^sub>C f), \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D (\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def D.whisker_right
by simp
also have "... = (\<eta> ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
- C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.interchange [of "\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c" "\<eta> ?c" "\<Phi>\<^sub>F (g, f)" "F g \<star>\<^sub>D F f"]
by simp
also have "... = (\<eta> ?c \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<cdot>\<^sub>D (F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
using f g fg D.comp_arr_dom D.comp_cod_arr
- C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.interchange [of "\<eta> ?c" "F\<^sub>0 ?c" "F (g \<star>\<^sub>C f)" "\<Phi>\<^sub>F (g, f)"]
by simp
also have "... = ((\<eta> ?c \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
- using f g fg ** C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right
+ using f g fg ** C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def D.whisker_right
by simp
finally have "(((\<tau>\<^sub>0' ?c \<star>\<^sub>D \<tau>\<^sub>0 ?c) \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F g \<star>\<^sub>D F f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= ((\<eta> ?c \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D ((F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F (g \<star>\<^sub>C f), \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "((F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg ** C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right
+ using f g fg ** C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def D.whisker_right
by simp
also have "... = \<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<cdot>\<^sub>D \<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def
D.lunit'_naturality [of "\<Phi>\<^sub>F (g, f)"]
by simp
also have "... = (\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
- using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right
+ using f g fg C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C F.FF_def D.whisker_right
by simp
finally have "((F\<^sub>0 ?c \<star>\<^sub>D \<Phi>\<^sub>F (g, f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<l>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D F f] \<star>\<^sub>D \<tau>\<^sub>0' ?a)
= (\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
(* One more associativity to move to its final position. *)
also have "... = (\<tau>\<^sub>0' ?c \<star>\<^sub>D \<r>\<^sub>D[G (g \<star>\<^sub>C f)]) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D G (g \<star>\<^sub>C f) \<star>\<^sub>D \<epsilon> ?a) \<cdot>\<^sub>D
(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<a>\<^sub>D[G (g \<star>\<^sub>C f), \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a]) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, G (g \<star>\<^sub>C f) \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c, F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
((\<eta> ?c \<star>\<^sub>D F (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<l>\<^sub>D\<^sup>-\<^sup>1[F (g \<star>\<^sub>C f)] \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
(\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
proof -
have "(\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f)) \<star>\<^sub>D \<tau>\<^sub>0' ?a) \<cdot>\<^sub>D
\<a>\<^sub>D[\<tau>\<^sub>0' ?c, \<tau>\<^sub>0 ?c \<star>\<^sub>D F (g \<star>\<^sub>C f), \<tau>\<^sub>0' ?a]
= \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G (g \<star>\<^sub>C f) \<star>\<^sub>D \<tau>\<^sub>0 ?a, \<tau>\<^sub>0' ?a] \<cdot>\<^sub>D
((\<tau>\<^sub>0' ?c \<star>\<^sub>D D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
using f g fg \<tau>.iso_map\<^sub>1_ide
D.assoc_naturality [of "\<tau>\<^sub>0' ?c" "D.inv (\<tau>\<^sub>1 (g \<star>\<^sub>C f))" "\<tau>\<^sub>0' ?a"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<tau>\<^sub>1' (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
unfolding map\<^sub>1_def
using f g fg D.comp_assoc by simp
finally show "(\<tau>\<^sub>0' ?c \<star>\<^sub>D \<Phi>\<^sub>G (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[\<tau>\<^sub>0' ?c, G g, G f] \<cdot>\<^sub>D (\<tau>\<^sub>1' g \<star>\<^sub>D G f) \<cdot>\<^sub>D
D.inv \<a>\<^sub>D[F g, \<tau>\<^sub>0' (src\<^sub>C g), G f] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<tau>\<^sub>1' f) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, \<tau>\<^sub>0' ?a]
= \<tau>\<^sub>1' (g \<star>\<^sub>C f) \<cdot>\<^sub>D (\<Phi>\<^sub>F (g, f) \<star>\<^sub>D \<tau>\<^sub>0' ?a)"
by blast
qed
qed
lemma is_pseudonatural_equivalence:
shows "pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G F \<Phi>\<^sub>F \<tau>\<^sub>0' \<tau>\<^sub>1'"
..
end
subsection "Pseudonaturally Equivalent Pseudofunctors"
text \<open>
Pseudofunctors \<open>F\<close> and \<open>G\<close> are \emph{pseudonaturally equivalent} if there is a
pseudonatural equivalence between them.
\<close>
definition pseudonaturally_equivalent
where "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<equiv>
\<exists>\<tau>\<^sub>0 \<tau>\<^sub>1. pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1"
lemma pseudonaturally_equivalent_reflexive:
assumes "pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F"
shows "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F F \<Phi>\<^sub>F"
proof -
interpret F: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F
using assms by simp
interpret identity_pseudonatural_transformation
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F
..
show ?thesis
using pseudonatural_equivalence_axioms pseudonaturally_equivalent_def by blast
qed
lemma pseudonaturally_equivalent_symmetric:
assumes "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G"
shows "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G F \<Phi>\<^sub>F"
proof -
obtain \<tau>\<^sub>0 \<tau>\<^sub>1 where \<tau>: "pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1"
using assms pseudonaturally_equivalent_def by blast
interpret \<tau>: pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
using \<tau> by simp
interpret \<sigma>: converse_pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
..
show ?thesis
using \<sigma>.is_pseudonatural_equivalence pseudonaturally_equivalent_def by blast
qed
lemma pseudonaturally_equivalent_transitive:
assumes "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G"
and "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G H \<Phi>\<^sub>H"
shows "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F H \<Phi>\<^sub>H"
proof -
obtain \<tau>\<^sub>0 \<tau>\<^sub>1 where \<tau>: "pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1"
using assms pseudonaturally_equivalent_def by blast
interpret \<tau>: pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
using \<tau> by simp
obtain \<sigma>\<^sub>0 \<sigma>\<^sub>1 where \<sigma>: "pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G H \<Phi>\<^sub>H \<sigma>\<^sub>0 \<sigma>\<^sub>1"
using assms pseudonaturally_equivalent_def by blast
interpret \<sigma>: pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G H \<Phi>\<^sub>H \<sigma>\<^sub>0 \<sigma>\<^sub>1
using \<sigma> by simp
interpret \<sigma>\<tau>: composite_pseudonatural_equivalence V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G H \<Phi>\<^sub>H \<tau>\<^sub>0 \<tau>\<^sub>1 \<sigma>\<^sub>0 \<sigma>\<^sub>1
..
show ?thesis
using \<sigma>\<tau>.pseudonatural_equivalence_axioms pseudonaturally_equivalent_def by blast
qed
text \<open>
If \<open>\<tau>\<close> is a pseudonatural equivalence from pseudofunctor \<open>F\<close> to pseudofunctor \<open>G\<close>
and \<open>\<sigma>\<close> is the converse equivalence, then \<open>F\<close> is locally naturally isomorphic to
the functor that takes a 2-cell \<open>\<mu>\<close> of \<open>C\<close> to \<open>\<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)\<close>
and symmetrically for \<open>G\<close>. Here we just establish the naturality property and
and that each 1-cell \<open>\<guillemotleft>g : a \<rightarrow>\<^sub>C a'\<guillemotright>\<close> is isomorphic to \<open>\<tau>\<^sub>0 a' \<star>\<^sub>D (\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a\<close>.
We ultimately need these results to prove that a pseudofunctor extends to an
equivalence of bicategories if and only if it is an equivalence pseudofunctor.
\<close>
context pseudonatural_equivalence
begin
interpretation \<sigma>: converse_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
..
abbreviation (input) \<sigma>\<^sub>0
where "\<sigma>\<^sub>0 \<equiv> \<sigma>.map\<^sub>0"
definition \<phi>
where "\<phi> f \<equiv> \<a>\<^sub>D[\<tau>\<^sub>0 (trg\<^sub>C f), F f, \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(\<tau>\<^sub>1 f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[G f, \<tau>\<^sub>0 (src\<^sub>C f), \<sigma>\<^sub>0 (src\<^sub>C f)] \<cdot>\<^sub>D
(G f \<star>\<^sub>D D.inv (\<sigma>.\<epsilon> (src\<^sub>C f))) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[G f]"
lemma \<phi>_in_hom [intro]:
assumes "C.ide f"
shows "\<guillemotleft>\<phi> f : G f \<Rightarrow>\<^sub>D \<tau>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D F f \<star>\<^sub>D \<sigma>\<^sub>0 (src\<^sub>C f)\<guillemotright>"
unfolding \<phi>_def
using assms by (intro D.comp_in_homI' D.hseqI') auto
lemma iso_\<phi>:
assumes "C.ide f"
shows "D.iso (\<phi> f)"
unfolding \<phi>_def
using assms iso_map\<^sub>1_ide
by (intro D.isos_compose) auto
definition \<psi>
where "\<psi> f \<equiv> (\<sigma>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D D.inv (\<tau>\<^sub>1 f)) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 (trg\<^sub>C f), \<tau>\<^sub>0 (trg\<^sub>C f), F f] \<cdot>\<^sub>D
(\<sigma>.\<eta> (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F f]"
lemma \<psi>_in_hom [intro]:
assumes "C.ide f"
shows "\<guillemotleft>\<psi> f : F f \<Rightarrow>\<^sub>D \<sigma>\<^sub>0 (trg\<^sub>C f) \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)\<guillemotright>"
unfolding \<psi>_def
using assms iso_map\<^sub>1_ide
by (intro D.comp_in_homI' D.hseqI') auto
lemma iso_\<psi>:
assumes "C.ide f"
shows "D.iso (\<psi> f)"
unfolding \<psi>_def
using assms iso_map\<^sub>1_ide
by (intro D.isos_compose) auto
lemma \<psi>_naturality:
assumes "C.arr \<mu>"
shows "(\<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D \<psi> (C.dom \<mu>) = \<psi> (C.cod \<mu>) \<cdot>\<^sub>D F \<mu>"
and "D.inv (\<psi> (C.cod \<mu>)) \<cdot>\<^sub>D (\<sigma>\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D \<psi> (C.dom \<mu>) = F \<mu>"
proof -
let ?a = "src\<^sub>C \<mu>" and ?a' = "trg\<^sub>C \<mu>"
let ?f = "C.dom \<mu>" and ?f' = "C.cod \<mu>"
have "(\<sigma>\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<psi> (C.dom \<mu>)
= ((\<sigma>\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D
(\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f))) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 ?a', \<tau>\<^sub>0 ?a', F ?f] \<cdot>\<^sub>D
(\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f]"
unfolding \<psi>_def
using assms D.comp_assoc by simp
also have "... = (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f')) \<cdot>\<^sub>D
((\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 ?a', \<tau>\<^sub>0 ?a', F ?f]) \<cdot>\<^sub>D
(\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f]"
proof -
have "(\<sigma>\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f))
= \<sigma>\<^sub>0 ?a' \<star>\<^sub>D (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)"
using assms D.whisker_left iso_map\<^sub>1_ide by simp
also have "... = \<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f') \<cdot>\<^sub>D (\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>)"
using assms naturality iso_map\<^sub>1_ide
D.invert_opposite_sides_of_square
[of "\<tau>\<^sub>1 ?f'" "G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a" "\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>" "\<tau>\<^sub>1 ?f"]
by simp
also have "... = (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f')) \<cdot>\<^sub>D (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>)"
using assms D.whisker_left iso_map\<^sub>1_ide by simp
finally have "(\<sigma>\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f))
= (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f')) \<cdot>\<^sub>D (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f')) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 ?a', \<tau>\<^sub>0 ?a', F ?f'] \<cdot>\<^sub>D
(((\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a') \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D
(\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f)) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f]"
proof -
have "(\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<a>\<^sub>D[\<sigma>\<^sub>0 ?a', \<tau>\<^sub>0 ?a', F ?f]
= \<a>\<^sub>D[\<sigma>\<^sub>0 ?a', \<tau>\<^sub>0 ?a', F ?f'] \<cdot>\<^sub>D ((\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a') \<star>\<^sub>D F \<mu>)"
using assms D.assoc_naturality [of "\<sigma>\<^sub>0 ?a'" "\<tau>\<^sub>0 ?a'" "F \<mu>"] by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D D.inv (\<tau>\<^sub>1 ?f')) \<cdot>\<^sub>D
\<a>\<^sub>D[\<sigma>\<^sub>0 ?a', \<tau>\<^sub>0 ?a', F ?f'] \<cdot>\<^sub>D
(\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f') \<cdot>\<^sub>D
(F.map\<^sub>0 ?a' \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D
\<l>\<^sub>D\<^sup>-\<^sup>1[F ?f]"
proof -
have "((\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a') \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D (\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f)
= (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a') \<cdot>\<^sub>D \<sigma>.\<eta> ?a' \<star>\<^sub>D F \<mu> \<cdot>\<^sub>D F ?f"
using assms D.interchange [of "\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a'" "\<sigma>.\<eta> ?a'" "F \<mu>" "F ?f"]
by simp
also have "... = \<sigma>.\<eta> ?a' \<cdot>\<^sub>D F.map\<^sub>0 ?a' \<star>\<^sub>D F ?f' \<cdot>\<^sub>D F \<mu>"
using assms D.comp_arr_dom D.comp_cod_arr by simp
also have "... = (\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f') \<cdot>\<^sub>D (F.map\<^sub>0 ?a' \<star>\<^sub>D F \<mu>)"
using assms D.interchange [of "\<sigma>.\<eta> ?a'" "F.map\<^sub>0 ?a'" "F ?f'" "F \<mu>"]
\<sigma>.unit_in_hom [of ?a']
by fastforce
finally have "((\<sigma>\<^sub>0 ?a' \<star>\<^sub>D \<tau>\<^sub>0 ?a') \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D (\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f)
= (\<sigma>.\<eta> ?a' \<star>\<^sub>D F ?f') \<cdot>\<^sub>D (F.map\<^sub>0 ?a' \<star>\<^sub>D F \<mu>)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<psi> ?f' \<cdot>\<^sub>D F \<mu>"
proof -
have "(F.map\<^sub>0 ?a' \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F ?f] = \<l>\<^sub>D\<^sup>-\<^sup>1[F ?f'] \<cdot>\<^sub>D F \<mu>"
using assms D.lunit'_naturality [of "F \<mu>"] by simp
thus ?thesis
unfolding \<psi>_def
using assms D.comp_assoc by simp
qed
finally show "(\<sigma>\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<psi> (C.dom \<mu>) = \<psi> ?f' \<cdot>\<^sub>D F \<mu>"
by blast
thus "D.inv (\<psi> ?f') \<cdot>\<^sub>D (\<sigma>\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<psi> (C.dom \<mu>) = F \<mu>"
using assms \<psi>_in_hom iso_\<psi>
by (metis C.ide_cod D.in_homE D.invert_side_of_triangle(1) D.seqI
F.preserves_arr F.preserves_cod)
qed
lemma isomorphic_expansion:
assumes "C.obj a" and "C.obj a'" and "\<guillemotleft>g : G.map\<^sub>0 a \<rightarrow>\<^sub>D G.map\<^sub>0 a'\<guillemotright>" and "D.ide g"
shows "D.isomorphic (\<tau>\<^sub>0 a' \<star>\<^sub>D (\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a) g"
proof -
let ?g' = "\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a"
have g': "\<guillemotleft>?g' : F.map\<^sub>0 a \<rightarrow>\<^sub>D F.map\<^sub>0 a'\<guillemotright>"
using assms ide_map\<^sub>0_obj \<sigma>.map\<^sub>0_simps(3) C.obj_simps
by (intro D.in_hhomI D.hseqI') auto
have ide_g': "D.ide ?g'"
using assms g' \<sigma>.ide_map\<^sub>0_obj ide_map\<^sub>0_obj by blast
let ?\<psi> = "(\<sigma>\<^sub>0 a' \<star>\<^sub>D \<r>\<^sub>D[g]) \<cdot>\<^sub>D (\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<sigma>.\<epsilon> a) \<cdot>\<^sub>D
(\<sigma>\<^sub>0 a' \<star>\<^sub>D \<a>\<^sub>D[g, \<tau>\<^sub>0 a, \<sigma>\<^sub>0 a]) \<cdot>\<^sub>D \<a>\<^sub>D[\<sigma>\<^sub>0 a', g \<star>\<^sub>D \<tau>\<^sub>0 a, \<sigma>\<^sub>0 a]"
have \<psi>: "\<guillemotleft>?\<psi> : ?g' \<star>\<^sub>D \<sigma>\<^sub>0 a \<Rightarrow>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D g\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<a>\<^sub>D[\<sigma>\<^sub>0 a', g \<star>\<^sub>D \<tau>\<^sub>0 a, \<sigma>\<^sub>0 a] :
(\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a \<Rightarrow>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D (g \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a\<guillemotright>"
using assms g' by fastforce
show "\<guillemotleft>\<sigma>\<^sub>0 a' \<star>\<^sub>D \<a>\<^sub>D[g, \<tau>\<^sub>0 a, \<sigma>\<^sub>0 a] :
\<sigma>\<^sub>0 a' \<star>\<^sub>D (g \<star>\<^sub>D \<tau>\<^sub>0 a) \<star>\<^sub>D \<sigma>\<^sub>0 a \<Rightarrow>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a\<guillemotright>"
using assms g' by fastforce
show "\<guillemotleft>\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<sigma>.\<epsilon> a :
\<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a \<star>\<^sub>D \<sigma>\<^sub>0 a \<Rightarrow>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D G.map\<^sub>0 a\<guillemotright>"
using assms C.obj_simps
by (intro D.hcomp_in_vhom) auto
show "\<guillemotleft>\<sigma>\<^sub>0 a' \<star>\<^sub>D \<r>\<^sub>D[g] : \<sigma>\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D G.map\<^sub>0 a \<Rightarrow>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D g\<guillemotright>"
using assms
apply (intro D.hcomp_in_vhom)
apply auto
by fastforce
qed
have iso_\<psi>: "D.iso ?\<psi>"
using assms g' \<psi> ide_g' ide_map\<^sub>0_obj C.obj_simps
by (intro D.isos_compose) auto
let ?\<phi> = "\<l>\<^sub>D[g] \<cdot>\<^sub>D (\<sigma>.\<epsilon> a' \<star>\<^sub>D g) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a', \<sigma>\<^sub>0 a', g] \<cdot>\<^sub>D (\<tau>\<^sub>0 a' \<star>\<^sub>D ?\<psi>)"
have \<phi>: "\<guillemotleft>?\<phi> : \<tau>\<^sub>0 a' \<star>\<^sub>D ?g' \<star>\<^sub>D \<sigma>\<^sub>0 a \<Rightarrow>\<^sub>D g\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<tau>\<^sub>0 a' \<star>\<^sub>D ?\<psi> : \<tau>\<^sub>0 a' \<star>\<^sub>D ?g' \<star>\<^sub>D \<sigma>\<^sub>0 a \<Rightarrow>\<^sub>D \<tau>\<^sub>0 a' \<star>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D g\<guillemotright>"
using assms g' \<psi> C.obj_simps
by (intro D.hcomp_in_vhom) auto
show "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[\<tau>\<^sub>0 a', \<sigma>\<^sub>0 a', g] : \<tau>\<^sub>0 a' \<star>\<^sub>D \<sigma>\<^sub>0 a' \<star>\<^sub>D g \<Rightarrow>\<^sub>D (\<tau>\<^sub>0 a' \<star>\<^sub>D \<sigma>\<^sub>0 a') \<star>\<^sub>D g\<guillemotright>"
proof -
have "src\<^sub>D (\<tau>\<^sub>0 a') = trg\<^sub>D (\<sigma>\<^sub>0 a')"
using assms by auto
moreover have "src\<^sub>D (\<sigma>\<^sub>0 a') = trg\<^sub>D g"
using assms by auto
ultimately show ?thesis
using assms ide_map\<^sub>0_obj \<sigma>.ide_map\<^sub>0_obj by auto
qed
show "\<guillemotleft>\<sigma>.\<epsilon> a' \<star>\<^sub>D g : (\<tau>\<^sub>0 a' \<star>\<^sub>D \<sigma>\<^sub>0 a') \<star>\<^sub>D g \<Rightarrow>\<^sub>D G.map\<^sub>0 a' \<star>\<^sub>D g\<guillemotright>"
using assms by fastforce
show "\<guillemotleft>\<l>\<^sub>D[g] : G.map\<^sub>0 a' \<star>\<^sub>D g \<Rightarrow>\<^sub>D g\<guillemotright>"
using assms by auto
qed
have iso_\<phi>: "D.iso ?\<phi>"
using assms g' \<phi> \<psi> iso_\<psi> ide_map\<^sub>0_obj iso_map\<^sub>1_ide \<sigma>.ide_map\<^sub>0_obj
apply (intro D.isos_compose)
apply (meson D.arrI D.hseqE D.ide_is_iso D.iso_hcomp D.seqE)
apply (metis D.hseqE D.in_hhomE D.iso_assoc' D.trg_hcomp F.map\<^sub>0_def map\<^sub>0_simps(2))
by auto
show ?thesis
using \<phi> iso_\<phi> D.isomorphic_def by auto
qed
end
text \<open>
Here we show that pseudonatural equivalence respects equivalence pseudofunctors,
in the sense that a pseudofunctor \<open>G\<close>, pseudonaturally equivalent to an
equivalence pseudofunctor \<open>F\<close>, is itself an equivalence pseudofunctor.
\<close>
locale pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F +
pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G +
\<tau>: pseudonatural_equivalence V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>F :: "'c * 'c \<Rightarrow> 'd"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
and \<tau>\<^sub>0 :: "'c \<Rightarrow> 'd"
and \<tau>\<^sub>1 :: "'c \<Rightarrow> 'd"
begin
interpretation \<sigma>': converse_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
..
sublocale G: equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
proof
show "\<And>\<mu> \<mu>'. \<lbrakk>C.par \<mu> \<mu>'; G \<mu> = G \<mu>'\<rbrakk> \<Longrightarrow> \<mu> = \<mu>'"
proof -
fix \<mu> \<mu>'
assume par: "C.par \<mu> \<mu>'" and eq: "G \<mu> = G \<mu>'"
let ?a = "src\<^sub>C \<mu>" and ?a' = "trg\<^sub>C \<mu>"
let ?f = "C.dom \<mu>" and ?f' = "C.cod \<mu>"
have "\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu> = (\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 ?f \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)"
using par \<tau>.ide_map\<^sub>0_obj \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_arr_inv'
\<tau>.map\<^sub>1_in_hom [of \<mu>]
by (metis C.dom_trg C.ide_dom C.obj_trg C.trg.preserves_dom D.whisker_left
F.as_nat_trans.is_natural_2 F.as_nat_trans.naturality F.preserves_arr
\<tau>.map\<^sub>1_simps(5))
also have "... = ((\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>) \<cdot>\<^sub>D \<tau>\<^sub>1 ?f) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)"
using D.comp_assoc by simp
also have "... = ((\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>') \<cdot>\<^sub>D \<tau>\<^sub>1 ?f) \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)"
using eq par \<tau>.naturality [of \<mu>] \<tau>.naturality [of \<mu>'] C.src_cod C.trg_cod by metis
also have 1: "... = (\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>') \<cdot>\<^sub>D \<tau>\<^sub>1 ?f \<cdot>\<^sub>D D.inv (\<tau>\<^sub>1 ?f)"
using D.comp_assoc by blast
also have "... = \<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>'"
using par \<tau>.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_arr_inv' \<tau>.map\<^sub>1_in_hom [of \<mu>']
by (metis 1 C.ide_dom C.trg_cod D.hseqE D.hseqI' D.hseq_char' D.seqE
F.preserves_arr F.preserves_trg calculation)
finally have "\<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu> = \<tau>\<^sub>0 ?a' \<star>\<^sub>D F \<mu>'"
by blast
hence "F \<mu> = F \<mu>'"
using par \<tau>.components_are_equivalences
D.equivalence_cancel_left [of "\<tau>\<^sub>0 ?a'" "F \<mu>" "F \<mu>'"]
by simp
thus "\<mu> = \<mu>'"
using par F.is_faithful by blast
qed
show "\<And>b. D.obj b \<Longrightarrow> \<exists>a. C.obj a \<and> D.equivalent_objects (map\<^sub>0 a) b"
proof -
fix b
assume b: "D.obj b"
obtain a where a: "C.obj a \<and> D.equivalent_objects (F.map\<^sub>0 a) b"
using b F.biessentially_surjective_on_objects by blast
have "D.equivalent_objects (F.map\<^sub>0 a) (map\<^sub>0 a)"
using a \<tau>.components_are_equivalences D.equivalent_objects_def
by (metis F.map\<^sub>0_def map\<^sub>0_def \<tau>.map\<^sub>0_in_hhom)
hence "D.equivalent_objects (map\<^sub>0 a) b"
using a D.equivalent_objects_symmetric D.equivalent_objects_transitive by blast
thus "\<exists>a. C.obj a \<and> D.equivalent_objects (map\<^sub>0 a) b"
using a by auto
qed
show "\<And>a a' g. \<lbrakk>C.obj a; C.obj a'; \<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>; D.ide g\<rbrakk>
\<Longrightarrow> \<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>C a'\<guillemotright> \<and> C.ide f \<and> D.isomorphic (G f) g"
proof -
fix a a' g
assume a: "C.obj a" and a': "C.obj a'"
assume g: "\<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>" and ide_g: "D.ide g"
interpret \<sigma>: converse_pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
..
interpret \<sigma>: pseudonatural_equivalence
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
G \<Phi>\<^sub>G F \<Phi>\<^sub>F \<sigma>'.map\<^sub>0 \<sigma>'.map\<^sub>1
using \<sigma>'.is_pseudonatural_equivalence by simp
let ?g' = "\<sigma>'.map\<^sub>0 a' \<star>\<^sub>D g \<star>\<^sub>D \<tau>\<^sub>0 a"
have g': "\<guillemotleft>?g' : F.map\<^sub>0 a \<rightarrow>\<^sub>D F.map\<^sub>0 a'\<guillemotright>"
using a a' g \<tau>.ide_map\<^sub>0_obj \<sigma>'.map\<^sub>0_simps(3) C.obj_simps
by (intro D.in_hhomI D.hseqI') auto
have ide_g': "D.ide ?g'"
using a a' g g' ide_g
by (meson D.hcomp_in_hhomE D.hseqE D.ide_hcomp \<sigma>'.ide_map\<^sub>0_obj \<tau>.ide_map\<^sub>0_obj)
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C a'\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) ?g'"
using a a' g g' ide_g' F.locally_essentially_surjective [of a a' ?g'] by auto
obtain \<xi> where \<xi>: "\<guillemotleft>\<xi> : F f \<Rightarrow>\<^sub>D ?g'\<guillemotright> \<and> D.iso \<xi>"
using f D.isomorphic_def by blast
have "D.isomorphic (G f) g"
proof -
have "D.isomorphic (G f) (\<tau>\<^sub>0 a' \<star>\<^sub>D F f \<star>\<^sub>D \<sigma>'.map\<^sub>0 a)"
using a a' f \<tau>.iso_\<phi> D.isomorphic_def by blast
also have "D.isomorphic ... (\<tau>\<^sub>0 a' \<star>\<^sub>D ?g' \<star>\<^sub>D \<sigma>'.map\<^sub>0 a)"
using a a' g' \<xi> D.isomorphic_def
by (metis D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide D.hseqE D.ideD(1)
D.isomorphic_implies_ide(2) \<sigma>'.ide_map\<^sub>0_obj calculation \<tau>.ide_map\<^sub>0_obj)
also have "D.isomorphic ... g"
using a a' g ide_g \<tau>.isomorphic_expansion by simp
finally show ?thesis by blast
qed
thus "\<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>C a'\<guillemotright> \<and> C.ide f \<and> D.isomorphic (G f) g"
using f by auto
qed
show "\<And>f f' \<nu>. \<lbrakk>C.ide f; C.ide f'; src\<^sub>C f = src\<^sub>C f'; trg\<^sub>C f = trg\<^sub>C f'; \<guillemotleft>\<nu> : G f \<Rightarrow>\<^sub>D G f'\<guillemotright>\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> G \<mu> = \<nu>"
proof -
fix f f' \<nu>
assume f: "C.ide f" and f': "C.ide f'"
and eq_src: "src\<^sub>C f = src\<^sub>C f'" and eq_trg: "trg\<^sub>C f = trg\<^sub>C f'"
and \<nu>: "\<guillemotleft>\<nu> : G f \<Rightarrow>\<^sub>D G f'\<guillemotright>"
let ?a = "src\<^sub>C f" and ?a' = "trg\<^sub>C f"
let ?\<nu>' = "D.inv (\<tau>.\<psi> f') \<cdot>\<^sub>D (\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<tau>.\<psi> f"
have \<nu>': "\<guillemotleft>?\<nu>' : F f \<Rightarrow>\<^sub>D F f'\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<tau>.\<psi> f : F f \<Rightarrow>\<^sub>D \<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a\<guillemotright>"
using f \<tau>.\<psi>_in_hom [of f] by blast
show "\<guillemotleft>\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a :
\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D G f \<star>\<^sub>D \<tau>\<^sub>0 ?a \<Rightarrow>\<^sub>D \<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D G f' \<star>\<^sub>D \<tau>\<^sub>0 ?a\<guillemotright>"
using f \<nu> by (intro D.hcomp_in_vhom) auto
show "\<guillemotleft>D.inv (\<tau>.\<psi> f') : \<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D G f' \<star>\<^sub>D \<tau>\<^sub>0 ?a \<Rightarrow>\<^sub>D F f'\<guillemotright>"
using f' \<tau>.iso_\<psi> eq_src eq_trg by auto
qed
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> F \<mu> = ?\<nu>'"
using f f' eq_src eq_trg \<nu>' F.locally_full [of f f' ?\<nu>'] by blast
have "G \<mu> = \<nu>"
proof -
have "D.inv (\<tau>.\<psi> f') \<cdot>\<^sub>D (\<sigma>'.map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D \<tau>.\<psi> f =
D.inv (\<tau>.\<psi> f') \<cdot>\<^sub>D (\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<tau>.\<psi> f"
using \<mu> \<tau>.\<psi>_naturality [of \<mu>] by auto
hence "\<sigma>'.map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>) = \<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a"
using f f' \<nu>' \<tau>.iso_\<psi>
D.iso_cancel_left
[of "D.inv (\<tau>.\<psi> f')" "(\<sigma>'.map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)) \<cdot>\<^sub>D \<tau>.\<psi> f"
"(\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a) \<cdot>\<^sub>D \<tau>.\<psi> f"]
D.iso_cancel_right
[of "\<tau>.\<psi> f" "\<sigma>'.map\<^sub>0 (trg\<^sub>C \<mu>) \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C \<mu>)" "\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a"]
by auto
hence "\<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a = \<sigma>'.map\<^sub>0 ?a' \<star>\<^sub>D \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a"
using \<mu> by auto
moreover have "D.par (G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f)) (\<nu> \<star>\<^sub>D \<tau>\<^sub>0 (src\<^sub>C f))"
using f \<mu> \<nu> \<nu>'
apply (intro conjI D.hseqI')
apply (auto simp add: D.vconn_implies_hpar(1))
apply fastforce
by (metis C.ideD(1) C.vconn_implies_hpar(1) D.hcomp_simps(4) D.hseqE D.hseqI'
D.seqE D.vconn_implies_hpar(1) preserves_arr preserves_cod preserves_src
weak_arrow_of_homs_axioms category.in_homE horizontal_homs_def
weak_arrow_of_homs_def)
ultimately have "G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a = \<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a"
using f f' \<mu> \<nu> \<tau>.ide_map\<^sub>0_obj \<sigma>'.ide_map\<^sub>0_obj
D.equivalence_cancel_left [of "\<sigma>'.map\<^sub>0 ?a'" "G \<mu> \<star>\<^sub>D \<tau>\<^sub>0 ?a" "\<nu> \<star>\<^sub>D \<tau>\<^sub>0 ?a"]
\<sigma>'.components_are_equivalences
by auto
moreover have "D.par (G \<mu>) \<nu>"
using f f' \<mu> \<nu> by fastforce
ultimately show ?thesis
using f f' \<mu> \<nu> \<tau>.ide_map\<^sub>0_obj
D.equivalence_cancel_right [of "\<tau>\<^sub>0 ?a" "G \<mu>" "\<nu>"] \<tau>.components_are_equivalences
by auto
qed
thus "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> G \<mu> = \<nu>"
using \<mu> by auto
qed
qed
lemma is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G"
..
end
lemma pseudonaturally_equivalent_respects_equivalence_pseudofunctor:
assumes "pseudonaturally_equivalent
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G"
and "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F"
shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G"
proof -
obtain \<tau>\<^sub>0 \<tau>\<^sub>1 where \<tau>: "pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1"
using assms pseudonaturally_equivalent_def by blast
interpret F: equivalence_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F
using assms by simp
interpret G: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G
using assms \<tau>
by (simp add: pseudonatural_equivalence_def pseudonatural_transformation_def)
interpret \<tau>: pseudonatural_equivalence
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
using \<tau> by simp
interpret G: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F G \<Phi>\<^sub>G \<tau>\<^sub>0 \<tau>\<^sub>1
..
show "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C G \<Phi>\<^sub>G"
using G.is_equivalence_pseudofunctor by simp
qed
end
diff --git a/thys/Bicategory/SpanBicategory.thy b/thys/Bicategory/SpanBicategory.thy
--- a/thys/Bicategory/SpanBicategory.thy
+++ b/thys/Bicategory/SpanBicategory.thy
@@ -1,4479 +1,4479 @@
(* Title: SpanBicategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Span Bicategories"
text \<open>
In this section we construct the bicategory \<open>Span(C)\<close>, where \<open>C\<close> is a category with pullbacks.
The $0$-cells of \<open>Span(C)\<close> are the objects of \<open>C\<close>, the $1$-cells of \<open>Span(C)\<close> are pairs
\<open>(f\<^sub>0, f\<^sub>1)\<close> of arrows of \<open>C\<close> having a common domain, and the $2$-cells of \<open>Span(C)\<close>
are ``arrows of spans''. An arrow of spans from \<open>(f\<^sub>0, f\<^sub>1)\<close> to \<open>(g\<^sub>0, g\<^sub>1)\<close> is
an arrow \<open>\<guillemotleft>u: dom f\<^sub>0 \<rightarrow> dom g\<^sub>0\<guillemotright>\<close> of \<open>C\<close>, such that \<open>g\<^sub>0 \<cdot> u = f\<^sub>0\<close> and \<open>g\<^sub>1 \<cdot> u = f\<^sub>1\<close>.
In the present development, a \emph{span} is formalized as a structure \<open>\<lparr>Leg0 = f\<^sub>0, Leg1 = f\<^sub>1\<rparr>\<close>,
where \<open>f\<^sub>0\<close> and \<open>f\<^sub>1\<close> are arrows of \<open>C\<close> with a common domain, which we call the \emph{apex} of
the span.
An \emph{arrow of spans} is formalized as a structure \<open>\<lparr>Chn = u, Dom = S, Cod = T\<rparr>\<close>,
where \<open>S\<close> and \<open>T\<close> are spans, and \<open>\<guillemotleft>u : S.apex \<rightarrow> T.apex\<guillemotright>\<close> satisfies \<open>Leg0 T \<cdot> u = Leg0 S\<close>
and \<open>Leg1 T \<cdot> u = Leg1 S\<close>. We refer to the arrow \<open>u\<close> as the \emph{chine} of the arrow of spans.
Arrows of spans inherit a composition from that of \<open>C\<close>; this is ``vertical composition''.
Spans may be composed via pullback in \<open>C\<close>; this ``horizontal composition'' extends to
arrows of spans, so that it is functorial with respect to vertical composition.
These two compositions determine a bicategory, as we shall show.
\<close>
theory SpanBicategory
imports Bicategory InternalAdjunction Category3.FreeCategory Category3.CategoryWithPullbacks
begin
subsection "Spans"
record 'a span_data =
Leg0 :: 'a
Leg1 :: 'a
locale span_in_category =
C: category +
fixes S :: "'a span_data" (structure)
assumes is_span: "C.span (Leg0 S) (Leg1 S)"
begin
abbreviation leg0
where "leg0 \<equiv> Leg0 S"
abbreviation leg1
where "leg1 \<equiv> Leg1 S"
abbreviation src
where "src \<equiv> C.cod leg0"
abbreviation trg
where "trg \<equiv> C.cod leg1"
definition apex
where "apex \<equiv> C.dom leg0"
lemma ide_apex [simp]:
shows "C.ide apex"
using is_span apex_def by simp
lemma leg_in_hom [intro]:
shows "\<guillemotleft>leg0 : apex \<rightarrow> src\<guillemotright>"
and "\<guillemotleft>leg1 : apex \<rightarrow> trg\<guillemotright>"
using is_span apex_def by auto
lemma leg_simps [simp]:
shows "C.arr leg0" and "C.dom leg0 = apex"
and "C.arr leg1" and "C.dom leg1 = apex"
using leg_in_hom by auto
end
record 'a arrow_of_spans_data =
Chn :: 'a
Dom :: "'a span_data"
Cod :: "'a span_data"
locale arrow_of_spans =
C: category C +
dom: span_in_category C \<open>Dom \<mu>\<close> +
cod: span_in_category C \<open>Cod \<mu>\<close>
for C :: "'a comp" (infixr "\<cdot>" 55)
and \<mu> :: "'a arrow_of_spans_data" (structure) +
assumes chine_in_hom [intro]: "\<guillemotleft>Chn \<mu> : dom.apex \<rightarrow> cod.apex\<guillemotright>"
and leg0_commutes [simp]: "cod.leg0 \<cdot> Chn \<mu> = dom.leg0"
and leg1_commutes [simp]: "cod.leg1 \<cdot> (Chn \<mu>) = dom.leg1"
begin
abbreviation chine
where "chine \<equiv> Chn \<mu>"
lemma chine_simps [simp]:
shows "C.arr chine" and "C.dom chine = dom.apex" and "C.cod chine = cod.apex"
using chine_in_hom by auto
lemma cod_src_eq_dom_src [simp]:
shows "cod.src = dom.src"
using dom.is_span cod.is_span
by (metis C.cod_comp leg0_commutes)
lemma cod_trg_eq_dom_trg [simp]:
shows "cod.trg = dom.trg"
using dom.is_span cod.is_span
by (metis C.cod_comp leg1_commutes)
abbreviation dsrc
where "dsrc \<equiv> dom.src"
abbreviation dtrg
where "dtrg \<equiv> dom.trg"
end
locale identity_arrow_of_spans =
arrow_of_spans +
assumes chine_is_identity [simp]: "C.ide (Chn \<mu>)"
begin
abbreviation apex
where "apex \<equiv> dom.apex"
abbreviation leg0
where "leg0 \<equiv> dom.leg0"
abbreviation leg1
where "leg1 \<equiv> dom.leg1"
lemma chine_eq_apex [simp]:
shows "chine = apex"
using chine_is_identity C.ideD(2) chine_simps(2) by presburger
lemma cod_simps [simp]:
shows "cod.apex = apex" and "cod.leg0 = leg0" and "cod.leg1 = leg1"
using chine_is_identity chine_simps(3) C.comp_arr_ide leg0_commutes leg1_commutes
by force+
end
subsection "The Vertical Category of Spans"
text \<open>
The following locale constructs the category of spans and arrows of spans in
an underlying category C, which is not yet assumed to have pullbacks.
The composition is vertical composition of arrows of spans, to which we will
later add horizontal composition to obtain a bicategory.
\<close>
locale span_vertical_category =
C: category
begin
abbreviation Null
where "Null \<equiv> \<lparr>Chn = C.null,
Dom = \<lparr>Leg0 = C.null, Leg1 = C.null\<rparr>,
Cod = \<lparr>Leg0 = C.null, Leg1 = C.null\<rparr>\<rparr>"
lemma not_arr_Null:
shows "\<not> arrow_of_spans C Null"
unfolding arrow_of_spans_def arrow_of_spans_axioms_def
by auto
text \<open>
Arrows of spans are composed simply by composing their chines.
\<close>
definition vcomp
where "vcomp \<nu> \<mu> \<equiv> if arrow_of_spans C \<mu> \<and> arrow_of_spans C \<nu> \<and> Dom \<nu> = Cod \<mu>
then \<lparr>Chn = Chn \<nu> \<cdot> Chn \<mu>, Dom = Dom \<mu>, Cod = Cod \<nu>\<rparr>
else Null"
notation vcomp (infixr "\<bullet>" 55)
(*
* TODO: The reason why the this and the subsequent category interpretation are declared
* as V: is that subsequently proved facts with the same names as the partial_magma and
* category locales silently override the latter, resulting in problems proving things.
* The presence of the extra "V" is only an issue up until a later sublocale declaration
* inherits everything from horizontal_homs. I wish I could say that I completely
* understood the inheritance and overriding rules for locales.
*)
interpretation V: partial_magma vcomp
using not_arr_Null vcomp_def
apply unfold_locales
by (metis (no_types, opaque_lifting))
lemma is_partial_magma:
shows "partial_magma vcomp"
..
lemma null_char:
shows "V.null = Null"
using V.null_def vcomp_def not_arr_Null
by (metis (no_types, lifting) V.comp_null(2))
text \<open>
Identities are arrows of spans whose chines are identities of C.
\<close>
lemma ide_char:
shows "V.ide \<mu> \<longleftrightarrow> arrow_of_spans C \<mu> \<and> C.ide (Chn \<mu>)"
proof
show "V.ide \<mu> \<Longrightarrow> arrow_of_spans C \<mu> \<and> C.ide (Chn \<mu>)"
proof
assume 0: "V.ide \<mu>"
have 1: "vcomp \<mu> \<mu> \<noteq> Null \<and> (\<forall>\<nu>. (\<nu> \<bullet> \<mu> \<noteq> Null \<longrightarrow> \<nu> \<bullet> \<mu> = \<nu>) \<and>
(\<mu> \<bullet> \<nu> \<noteq> Null \<longrightarrow> \<mu> \<bullet> \<nu> = \<nu>))"
using 0 V.ide_def null_char by simp
show \<mu>: "arrow_of_spans C \<mu>"
using 1 vcomp_def by metis
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu> by auto
show "C.ide (Chn \<mu>)"
proof -
have "\<mu>.chine \<cdot> \<mu>.chine \<noteq> C.null"
using 1 vcomp_def
by (metis C.in_homE C.not_arr_null C.seqI \<mu>.chine_in_hom)
moreover have "\<And>f. f \<cdot> Chn \<mu> \<noteq> C.null \<Longrightarrow> f \<cdot> Chn \<mu> = f"
proof -
fix f
assume "f \<cdot> \<mu>.chine \<noteq> C.null"
hence f: "\<guillemotleft>f : \<mu>.cod.apex \<rightarrow> C.cod f\<guillemotright>"
using C.ext C.in_homI by force
let ?cod_\<mu> = "\<lparr>Chn = C.cod \<mu>.chine, Dom = Cod \<mu>, Cod = Cod \<mu>\<rparr>"
interpret cod_\<mu>: arrow_of_spans C ?cod_\<mu>
using C.ide_in_hom \<mu>.cod.ide_apex \<mu>.chine_in_hom C.comp_arr_dom
by (unfold_locales, auto)
have "?cod_\<mu> \<bullet> \<mu> = ?cod_\<mu>"
by (metis (no_types, opaque_lifting) "1" C.not_arr_null
\<mu>.cod.span_in_category_axioms arrow_of_spans_data.select_convs(2)
cod_\<mu>.arrow_of_spans_axioms span_data.select_convs(1)
span_in_category.leg_simps(1) vcomp_def)
thus "f \<cdot> \<mu>.chine = f"
using C.comp_arr_ide C.comp_cod_arr \<mu> cod_\<mu>.arrow_of_spans_axioms f vcomp_def
by auto
qed
moreover have "\<And>f. \<mu>.chine \<cdot> f \<noteq> C.null \<Longrightarrow> \<mu>.chine \<cdot> f = f"
by (metis C.comp_cod_arr C.comp_ide_arr C.ext C.ide_char' calculation(1-2))
ultimately show "C.ide \<mu>.chine"
unfolding C.ide_def by simp
qed
qed
show "arrow_of_spans C \<mu> \<and> C.ide (Chn \<mu>) \<Longrightarrow> V.ide \<mu>"
proof -
assume \<mu>: "arrow_of_spans C \<mu> \<and> C.ide (Chn \<mu>)"
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu> by auto
have 1: "Dom \<mu> = Cod \<mu>"
using \<mu> identity_arrow_of_spans.cod_simps(2) identity_arrow_of_spans.cod_simps(3)
identity_arrow_of_spans.intro identity_arrow_of_spans_axioms.intro
by fastforce
show "V.ide \<mu>"
proof -
have "\<mu> \<bullet> \<mu> \<noteq> V.null"
using \<mu> 1 vcomp_def by (simp add: C.ide_def null_char)
moreover have "\<And>\<nu>. vcomp \<nu> \<mu> \<noteq> V.null \<Longrightarrow> vcomp \<nu> \<mu> = \<nu>"
proof -
fix \<nu> :: "'a arrow_of_spans_data"
assume \<nu>: "\<nu> \<bullet> \<mu> \<noteq> V.null"
have 2: "arrow_of_spans C \<nu> \<and> Dom \<nu> = Cod \<mu>"
using \<nu> 1 vcomp_def by (metis V.comp_null(2))
interpret \<nu>: arrow_of_spans C \<nu>
using 2 by auto
show "\<nu> \<bullet> \<mu> = \<nu>"
unfolding vcomp_def
using \<mu> 1 2 C.comp_arr_ide by simp
qed
moreover have "\<And>\<nu>. \<mu> \<bullet> \<nu> \<noteq> V.null \<Longrightarrow> \<mu> \<bullet> \<nu> = \<nu>"
proof -
fix \<nu> :: "'a arrow_of_spans_data"
assume \<nu>: "\<mu> \<bullet> \<nu> \<noteq> V.null"
have 2: "arrow_of_spans C \<nu> \<and> Dom \<mu> = Cod \<nu>"
using \<nu> 1 vcomp_def by (metis V.comp_null(1))
interpret \<nu>: arrow_of_spans C \<nu>
using 2 by auto
show "\<mu> \<bullet> \<nu> = \<nu>"
unfolding vcomp_def
using \<mu> 1 2 C.comp_ide_arr by simp
qed
ultimately show ?thesis
unfolding V.ide_def by blast
qed
qed
qed
lemma has_domain_char:
shows "V.domains \<mu> \<noteq> {} \<longleftrightarrow> arrow_of_spans C \<mu>"
proof
show "V.domains \<mu> \<noteq> {} \<Longrightarrow> arrow_of_spans C \<mu>"
using V.domains_def null_char vcomp_def by fastforce
show "arrow_of_spans C \<mu> \<Longrightarrow> V.domains \<mu> \<noteq> {}"
proof -
assume \<mu>: "arrow_of_spans C \<mu>"
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu> by auto
let ?dom_\<mu> = "\<lparr>Chn = \<mu>.dom.apex, Dom = Dom \<mu>, Cod = Dom \<mu>\<rparr>"
interpret dom_\<mu>: arrow_of_spans C ?dom_\<mu>
using C.comp_arr_dom by (unfold_locales, auto)
have "?dom_\<mu> \<in> V.domains \<mu>"
proof -
have "V.ide ?dom_\<mu>"
using ide_char dom_\<mu>.arrow_of_spans_axioms by simp
moreover have "\<mu> \<bullet> ?dom_\<mu> \<noteq> V.null"
using \<mu> vcomp_def \<mu>.cod.span_in_category_axioms dom_\<mu>.arrow_of_spans_axioms
null_char span_in_category.leg_simps(1)
by fastforce
ultimately show ?thesis
unfolding V.domains_def by blast
qed
thus "V.domains \<mu> \<noteq> {}" by blast
qed
qed
lemma has_codomain_char:
shows "V.codomains \<mu> \<noteq> {} \<longleftrightarrow> arrow_of_spans C \<mu>"
proof
show "V.codomains \<mu> \<noteq> {} \<Longrightarrow> arrow_of_spans C \<mu>"
using V.codomains_def null_char vcomp_def by fastforce
show "arrow_of_spans C \<mu> \<Longrightarrow> V.codomains \<mu> \<noteq> {}"
proof -
assume \<mu>: "arrow_of_spans C \<mu>"
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu> by auto
let ?cod_f = "\<lparr>Chn = \<mu>.cod.apex, Dom = Cod \<mu>, Cod = Cod \<mu>\<rparr>"
interpret cod_f: arrow_of_spans C ?cod_f
using C.comp_arr_dom by (unfold_locales, auto)
have "?cod_f \<in> V.codomains \<mu>"
proof -
have "V.ide ?cod_f"
using ide_char cod_f.arrow_of_spans_axioms by simp
moreover have "?cod_f \<bullet> \<mu> \<noteq> V.null"
using \<mu> vcomp_def \<mu>.cod.span_in_category_axioms cod_f.arrow_of_spans_axioms
null_char span_in_category.leg_simps(1)
by fastforce
ultimately show ?thesis
unfolding V.codomains_def by blast
qed
thus "V.codomains \<mu> \<noteq> {}" by blast
qed
qed
lemma arr_char:
shows "V.arr \<mu> \<longleftrightarrow> arrow_of_spans C \<mu>"
unfolding V.arr_def
using has_domain_char has_codomain_char by simp
lemma seq_char:
shows "V.seq \<nu> \<mu> \<longleftrightarrow> arrow_of_spans C \<mu> \<and> arrow_of_spans C \<nu> \<and> Dom \<nu> = Cod \<mu>"
proof
show "V.seq \<nu> \<mu> \<Longrightarrow> arrow_of_spans C \<mu> \<and> arrow_of_spans C \<nu> \<and> Dom \<nu> = Cod \<mu>"
using vcomp_def by (metis V.not_arr_null null_char)
show "arrow_of_spans C \<mu> \<and> arrow_of_spans C \<nu> \<and> Dom \<nu> = Cod \<mu> \<Longrightarrow> V.seq \<nu> \<mu>"
proof -
assume 1: "arrow_of_spans C \<mu> \<and> arrow_of_spans C \<nu> \<and> Dom \<nu> = Cod \<mu>"
interpret \<mu>: arrow_of_spans C \<mu>
using 1 by auto
interpret \<nu>: arrow_of_spans C \<nu>
using 1 by auto
show "V.seq \<nu> \<mu>"
proof -
let ?\<nu>\<mu> = "\<lparr>Chn = Chn \<nu> \<cdot> Chn \<mu>, Dom = Dom \<mu>, Cod = Cod \<nu>\<rparr>"
have "\<nu> \<bullet> \<mu> = ?\<nu>\<mu>"
using 1 vcomp_def by metis
moreover have "V.arr ?\<nu>\<mu>"
proof -
interpret Dom: span_in_category C \<open>Dom ?\<nu>\<mu>\<close>
by (simp add: \<mu>.dom.span_in_category_axioms)
interpret Cod: span_in_category C \<open>Cod ?\<nu>\<mu>\<close>
by (simp add: \<nu>.cod.span_in_category_axioms)
have "arrow_of_spans C ?\<nu>\<mu>"
using 1 \<mu>.chine_in_hom \<nu>.chine_in_hom C.comp_reduce
by (unfold_locales, cases ?\<nu>\<mu>, auto)
thus ?thesis
using arr_char by blast
qed
ultimately show ?thesis by simp
qed
qed
qed
interpretation V: category vcomp
proof
show "\<And>\<mu>. (V.domains \<mu> \<noteq> {}) = (V.codomains \<mu> \<noteq> {})"
using has_domain_char has_codomain_char by simp
show "\<And>\<nu> \<mu>. \<nu> \<bullet> \<mu> \<noteq> V.null \<Longrightarrow> V.seq \<nu> \<mu>"
using seq_char vcomp_def null_char by metis
show "\<And>\<pi> \<nu> \<mu>. V.seq \<pi> \<nu> \<Longrightarrow> V.seq (\<pi> \<bullet> \<nu>) \<mu> \<Longrightarrow> V.seq \<nu> \<mu>"
using seq_char vcomp_def by (metis arrow_of_spans_data.select_convs(2))
show "\<And>\<pi> \<nu> \<mu>. V.seq \<pi> (\<nu> \<bullet> \<mu>) \<Longrightarrow> V.seq \<nu> \<mu> \<Longrightarrow> V.seq \<pi> \<nu>"
using seq_char vcomp_def by (metis arrow_of_spans_data.select_convs(3))
show "\<And>\<nu> \<mu> \<pi>. V.seq \<nu> \<mu> \<Longrightarrow> V.seq \<pi> \<nu> \<Longrightarrow> V.seq (\<pi> \<bullet> \<nu>) \<mu>"
using seq_char vcomp_def by (metis arr_char arrow_of_spans_data.select_convs(2))
show "\<And>\<nu> \<mu> \<pi>. V.seq \<nu> \<mu> \<Longrightarrow> V.seq \<pi> \<nu> \<Longrightarrow> (\<pi> \<bullet> \<nu>) \<bullet> \<mu> = \<pi> \<bullet> \<nu> \<bullet> \<mu>"
proof -
fix \<mu> \<nu> \<pi>
assume \<mu>\<nu>: "V.seq \<nu> \<mu>" and \<nu>\<pi>: "V.seq \<pi> \<nu>"
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu>\<nu> seq_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using \<mu>\<nu> seq_char by auto
interpret \<pi>: arrow_of_spans C \<pi>
using \<nu>\<pi> seq_char by auto
show "(\<pi> \<bullet> \<nu>) \<bullet> \<mu> = \<pi> \<bullet> \<nu> \<bullet> \<mu>"
unfolding vcomp_def
using \<mu>\<nu> \<nu>\<pi> seq_char \<mu>.chine_in_hom \<nu>.chine_in_hom \<pi>.chine_in_hom
by (simp add: C.comp_assoc, metis arr_char vcomp_def)
qed
qed
lemma is_category:
shows "category vcomp"
..
lemma dom_char:
shows "V.dom = (\<lambda>\<mu>. if V.arr \<mu> then
\<lparr>Chn = span_in_category.apex C (Dom \<mu>), Dom = Dom \<mu>, Cod = Dom \<mu>\<rparr>
else V.null)"
proof
fix \<mu>
have "\<not> V.arr \<mu> \<Longrightarrow> V.dom \<mu> = V.null"
by (simp add: V.arr_def V.dom_def)
moreover have "V.arr \<mu> \<Longrightarrow> V.dom \<mu> = \<lparr>Chn = span_in_category.apex C (Dom \<mu>),
Dom = Dom \<mu>, Cod = Dom \<mu>\<rparr>"
by (metis V.comp_arr_dom V.comp_ide_self V.ideD(1) V.ide_dom arrow_of_spans_data.cases
arrow_of_spans_data.select_convs(1-3) ide_char identity_arrow_of_spans.chine_eq_apex
identity_arrow_of_spans_axioms.intro identity_arrow_of_spans_def seq_char)
ultimately show "V.dom \<mu> = (if V.arr \<mu> then
\<lparr>Chn = span_in_category.apex C (Dom \<mu>),
Dom = Dom \<mu>, Cod = Dom \<mu>\<rparr>
else V.null)"
by argo
qed
lemma cod_char:
shows "V.cod = (\<lambda>\<mu>. if V.arr \<mu> then
\<lparr>Chn = span_in_category.apex C (Cod \<mu>), Dom = Cod \<mu>, Cod = Cod \<mu>\<rparr>
else V.null)"
proof
fix \<mu>
have "\<not> V.arr \<mu> \<Longrightarrow> V.cod \<mu> = V.null"
by (simp add: V.arr_def V.cod_def)
moreover have "V.arr \<mu> \<Longrightarrow> V.cod \<mu> = \<lparr>Chn = span_in_category.apex C (Cod \<mu>),
Dom = Cod \<mu>, Cod = Cod \<mu>\<rparr>"
by (metis V.arr_cod V.comp_cod_arr V.dom_cod dom_char span_vertical_category.seq_char
span_vertical_category_axioms)
ultimately show "V.cod \<mu> = (if V.arr \<mu> then
\<lparr>Chn = span_in_category.apex C (Cod \<mu>),
Dom = Cod \<mu>, Cod = Cod \<mu>\<rparr>
else V.null)"
by argo
qed
lemma vcomp_char:
shows "vcomp = (\<lambda>\<nu> \<mu>. if V.seq \<nu> \<mu> then
\<lparr>Chn = Chn \<nu> \<cdot> Chn \<mu>, Dom = Dom \<mu>, Cod = Cod \<nu>\<rparr>
else V.null)"
by (meson V.ext seq_char vcomp_def)
lemma vcomp_eq:
assumes "V.seq \<nu> \<mu>"
shows "\<nu> \<bullet> \<mu> = \<lparr>Chn = Chn \<nu> \<cdot> Chn \<mu>, Dom = Dom \<mu>, Cod = Cod \<nu>\<rparr>"
using assms vcomp_char by meson
lemma Chn_vcomp:
assumes "V.seq \<nu> \<mu>"
shows "Chn (\<nu> \<bullet> \<mu>) = Chn \<nu> \<cdot> Chn \<mu>"
using assms vcomp_eq [of \<nu> \<mu>] by simp
lemma ide_char':
shows "V.ide \<mu> \<longleftrightarrow> identity_arrow_of_spans C \<mu>"
using arr_char ide_char identity_arrow_of_spans_axioms_def identity_arrow_of_spans_def
identity_arrow_of_spans.axioms(1) identity_arrow_of_spans.chine_is_identity
by metis
lemma Chn_in_hom:
assumes "V.in_hom \<tau> f g"
shows "C.in_hom (Chn \<tau>) (Chn f) (Chn g)"
by (metis arr_char arrow_of_spans.chine_in_hom arrow_of_spans_data.select_convs(1)
assms category.in_homE is_category span_vertical_category.cod_char
span_vertical_category.dom_char span_vertical_category_axioms)
abbreviation mkIde
where "mkIde f0 f1 \<equiv>
\<lparr>Chn = C.dom f0, Dom = \<lparr>Leg0 = f0, Leg1 = f1\<rparr>, Cod = \<lparr>Leg0 = f0, Leg1 = f1\<rparr>\<rparr>"
lemma ide_mkIde:
assumes "C.span f0 f1"
shows "V.ide (mkIde f0 f1)"
proof -
interpret f: span_in_category C \<open>\<lparr>Leg0 = f0, Leg1 = f1\<rparr>\<close>
using assms by (unfold_locales, auto)
interpret ff: arrow_of_spans C \<open>mkIde f0 f1\<close>
using assms f.apex_def C.comp_arr_dom
by (unfold_locales, auto)
show ?thesis
using assms ff.arrow_of_spans_axioms ide_char by simp
qed
abbreviation mkObj
where "mkObj a \<equiv> mkIde a a"
lemma ide_mkObj:
assumes "C.ide a"
shows "V.ide (mkObj a)"
using assms ide_mkIde [of a a] by auto
lemma inverse_arrows:
assumes "V.arr \<mu>" and "C.iso (Chn \<mu>)"
shows "V.inverse_arrows \<mu> \<lparr>Chn = C.inv (Chn \<mu>), Dom = Cod \<mu>, Cod = Dom \<mu>\<rparr>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
let ?\<nu> = "\<lparr>Chn = C.inv (Chn \<mu>), Dom = Cod \<mu>, Cod = Dom \<mu>\<rparr>"
interpret \<nu>: arrow_of_spans C ?\<nu>
using assms C.invert_side_of_triangle(2) [of \<mu>.dom.leg0 \<mu>.cod.leg0 \<mu>.chine]
C.invert_side_of_triangle(2) [of \<mu>.dom.leg1 \<mu>.cod.leg1 \<mu>.chine]
by (unfold_locales, auto)
show "V.inverse_arrows \<mu> ?\<nu>"
proof
show "V.ide (?\<nu> \<bullet> \<mu>)"
by (metis C.invert_side_of_triangle(1) Chn_vcomp \<mu>.arrow_of_spans_axioms
\<mu>.chine_simps(1) \<mu>.chine_simps(2) \<mu>.dom.ide_apex \<nu>.arrow_of_spans_axioms
arr_char select_convs(1-2) assms(2) C.comp_arr_dom ide_char seq_char)
thus "V.ide (\<mu> \<bullet> ?\<nu>)"
by (metis C.comp_inv_arr' C.inv_inv C.iso_inv_iso V.ide_compE \<mu>.cod.ide_apex
\<nu>.chine_simps(2) arr_char select_convs(1-2) assms(2) ide_char
Chn_vcomp)
qed
qed
lemma iso_char:
shows "V.iso \<mu> \<longleftrightarrow> V.arr \<mu> \<and> C.iso (Chn \<mu>)"
proof
show "V.iso \<mu> \<Longrightarrow> V.arr \<mu> \<and> C.iso (Chn \<mu>)"
using vcomp_eq ide_char
by (metis C.iso_iff_section_and_retraction C.retractionI C.sectionI Chn_vcomp
V.arr_cod V.arr_dom V.comp_arr_inv' V.comp_inv_arr' V.ide_cod V.ide_dom
V.iso_is_arr)
show "V.arr \<mu> \<and> C.iso (Chn \<mu>) \<Longrightarrow> V.iso \<mu>"
using inverse_arrows by auto
qed
lemma inv_eq:
assumes "V.iso \<mu>"
shows "V.inv \<mu> = \<lparr>Chn = C.inv (Chn \<mu>), Dom = Cod \<mu>, Cod = Dom \<mu>\<rparr>"
using assms inverse_arrows iso_char by (simp add: V.inverse_unique)
end
subsection "Putting Spans in Homs"
context span_vertical_category
begin
interpretation V: category vcomp
using is_category by simp
definition src
where "src \<mu> \<equiv> if V.arr \<mu> then mkObj (C.cod (Leg0 (Dom \<mu>))) else V.null"
lemma ide_src [simp]:
assumes "V.arr \<mu>"
shows "V.ide (src \<mu>)"
using assms src_def arr_char ide_mkObj C.ide_cod
by (simp add: arrow_of_spans_def span_in_category.leg_simps(1))
interpretation src: endofunctor vcomp src
proof
show "\<And>\<mu>. \<not> V.arr \<mu> \<Longrightarrow> src \<mu> = V.null"
using arr_char by (simp add: src_def null_char)
show 1: "\<And>\<mu>. V.arr \<mu> \<Longrightarrow> V.arr (src \<mu>)"
using ide_src by simp
show 2: "\<And>\<mu>. V.arr \<mu> \<Longrightarrow> V.dom (src \<mu>) = src (V.dom \<mu>)"
using 1 arr_char src_def dom_char ide_src V.arr_dom V.ideD(2) by force
show 3: "\<And>\<mu>. V.arr \<mu> \<Longrightarrow> V.cod (src \<mu>) = src (V.cod \<mu>)"
using 1 arr_char src_def cod_char ide_src V.arr_cod V.ideD(3)
arrow_of_spans.cod_src_eq_dom_src
by force
show "\<And>\<mu> \<nu>. V.seq \<nu> \<mu> \<Longrightarrow> src (\<nu> \<bullet> \<mu>) = src \<nu> \<bullet> src \<mu>"
by (metis (no_types, lifting) "1" "2" "3" V.comp_ide_self V.dom_comp V.ideD(2)
V.seqE span_vertical_category.ide_src span_vertical_category_axioms)
qed
lemma src_is_endofunctor:
shows "endofunctor vcomp src"
..
lemma src_vcomp:
assumes "V.seq \<nu> \<mu>"
shows "src (\<nu> \<bullet> \<mu>) = src \<nu> \<bullet> src \<mu>"
using assms src.preserves_comp by simp
definition trg
where "trg \<mu> \<equiv> if V.arr \<mu> then mkObj (C.cod (Leg1 (Dom \<mu>))) else V.null"
lemma ide_trg [simp]:
assumes "V.arr \<mu>"
shows "V.ide (trg \<mu>)"
using assms trg_def arr_char ide_mkObj C.ide_cod
by (simp add: arrow_of_spans_def span_in_category.leg_simps(3))
interpretation trg: endofunctor vcomp trg
proof
show "\<And>\<mu>. \<not> V.arr \<mu> \<Longrightarrow> trg \<mu> = V.null"
using arr_char by (simp add: trg_def null_char)
show 1: "\<And>\<mu>. V.arr \<mu> \<Longrightarrow> V.arr (trg \<mu>)"
using ide_trg by simp
show 2: "\<And>\<mu>. V.arr \<mu> \<Longrightarrow> V.dom (trg \<mu>) = trg (V.dom \<mu>)"
using 1 arr_char trg_def dom_char ide_trg V.arr_dom V.ideD(2) by force
show 3: "\<And>\<mu>. V.arr \<mu> \<Longrightarrow> V.cod (trg \<mu>) = trg (V.cod \<mu>)"
using 1 arr_char trg_def cod_char ide_trg V.arr_cod V.ideD(3)
arrow_of_spans.cod_trg_eq_dom_trg
by force
show "\<And>\<mu> \<nu>. V.seq \<nu> \<mu> \<Longrightarrow> trg (\<nu> \<bullet> \<mu>) = trg \<nu> \<bullet> trg \<mu>"
by (metis "2" "3" V.comp_ide_self V.dom_comp V.ide_char V.seqE ide_trg)
qed
lemma trg_is_endofunctor:
shows "endofunctor vcomp trg"
..
lemma trg_vcomp:
assumes "V.seq \<nu> \<mu>"
shows "trg (\<nu> \<bullet> \<mu>) = trg \<nu> \<bullet> trg \<mu>"
using assms trg.preserves_comp by simp
lemma src_trg_simps [simp]:
assumes "V.arr \<mu>"
shows "src (src \<mu>) = src \<mu>"
and "src (trg \<mu>) = trg \<mu>"
and "trg (src \<mu>) = src \<mu>"
and "trg (trg \<mu>) = trg \<mu>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
have 1: "V.arr \<lparr>Chn = \<mu>.dsrc, Dom = \<lparr>Leg0 = \<mu>.dsrc, Leg1 = \<mu>.dsrc\<rparr>,
Cod = \<lparr>Leg0 = \<mu>.dsrc, Leg1 = \<mu>.dsrc\<rparr>\<rparr>"
using ide_mkObj by auto
have 2: "V.arr \<lparr>Chn = \<mu>.dtrg, Dom = \<lparr>Leg0 = \<mu>.dtrg, Leg1 = \<mu>.dtrg\<rparr>,
Cod = \<lparr>Leg0 = \<mu>.dtrg, Leg1 = \<mu>.dtrg\<rparr>\<rparr>"
using ide_mkObj by auto
show "src (src \<mu>) = src \<mu>"
using assms 1 src_def by simp
show "trg (src \<mu>) = src \<mu>"
using assms 1 src_def trg_def by simp
show "src (trg \<mu>) = trg \<mu>"
using assms 2 src_def trg_def by simp
show "trg (trg \<mu>) = trg \<mu>"
using assms 2 trg_def by simp
qed
sublocale horizontal_homs vcomp src trg
by (unfold_locales, simp_all)
lemma has_horizontal_homs:
shows "horizontal_homs vcomp src trg"
..
lemma obj_char:
shows "obj a \<longleftrightarrow> V.ide a \<and> a = mkObj (Chn a)"
by (metis C.dom_cod V.comp_ide_self V.ide_char arrow_of_spans.chine_simps(3)
arrow_of_spans_data.select_convs(1,3) objE objI_trg span_data.select_convs(2)
cod_char seq_char trg_def)
end
subsection "Horizontal Composite of Spans"
text \<open>
We now define the horizontal composite \<open>S \<star> T\<close> of spans \<open>S\<close> and \<open>T\<close>,
assuming that \<open>C\<close> is a category with chosen pullbacks.
We think of Leg0 as an input and Leg1 as an output.
The following then defines the composite span \<open>S \<star> T\<close>, with \<open>T\<close> on the ``input side'' of \<open>S\<close>.
The notation is such that the \<open>\<p>\<^sub>0\<close> projections of \<open>C\<close> are used for legs on the input
(\emph{i.e.} the ``0'') side and the \<open>\<p>\<^sub>1\<close> projections are used for legs on the output
(\emph{i.e.} the ``1'') side.
\<close>
locale composite_span =
C: elementary_category_with_pullbacks +
S: span_in_category C S +
T: span_in_category C T
for S (structure)
and T (structure) +
assumes composable: "C.cod (Leg0 S) = C.cod (Leg1 T)"
begin
abbreviation this
where "this \<equiv> \<lparr>Leg0 = T.leg0 \<cdot> \<p>\<^sub>0[S.leg0, T.leg1], Leg1 = S.leg1 \<cdot> \<p>\<^sub>1[S.leg0, T.leg1]\<rparr>"
lemma leg0_prj_in_hom:
shows "\<guillemotleft>T.leg0 \<cdot> \<p>\<^sub>0[S.leg0, T.leg1] : S.leg0 \<down>\<down> T.leg1 \<rightarrow> C.cod (Leg0 T)\<guillemotright>"
using S.is_span T.is_span C.prj0_in_hom [of "Leg0 S" "Leg1 T"] composable by auto
lemma leg1_prj_in_hom:
shows "\<guillemotleft>S.leg1 \<cdot> \<p>\<^sub>1[S.leg0, T.leg1] : S.leg0 \<down>\<down> T.leg1 \<rightarrow> C.cod (Leg1 S)\<guillemotright>"
using S.is_span T.is_span C.prj1_in_hom [of "Leg0 S" "Leg1 T"] composable by auto
lemma is_span [simp]:
shows "span_in_category C this"
using leg0_prj_in_hom leg1_prj_in_hom
by (unfold_locales, fastforce)
sublocale span_in_category C this
using is_span by auto
end
locale span_bicategory =
C: elementary_category_with_pullbacks +
span_vertical_category
begin
definition chine_hcomp
where "chine_hcomp \<nu> \<mu> \<equiv>
\<langle>Chn \<nu> \<cdot> \<p>\<^sub>1[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)]
\<lbrakk>Leg0 (Cod \<nu>), Leg1 (Cod \<mu>)\<rbrakk>
Chn \<mu> \<cdot> \<p>\<^sub>0[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)]\<rangle>"
text \<open>
$$\xymatrix{
& & \scriptstyle{{\rm src}({\rm Dom}~\nu)} \;=\; {{\rm trg}({\rm Dom}~\mu)} & &\\
&
\ar[ddl] _{{\rm Leg1}({\rm Dom}~\nu)}
\ar [ur] ^<>(0.4){{\rm Leg0}({\rm Dom}~\nu)\hspace{20pt}}
\ar[dddd] ^{{\rm Chn}~\nu}
&
&
\ar[ul] _<>(0.4){\hspace{20pt}{\rm Leg1}({\rm Dom}~\mu)}
\ar[ddr] ^{{\rm Leg0}({\rm Dom}~\mu)}
\ar[dddd] _{{\rm Chn}~\mu}
\\
& &
\ar[ul] ^{p_1}
\ar[ur] _{p_0}
\ar@ {.>}[dd]^<>(0.3){{\rm chn\_hcomp~\mu~\nu}}
\\
\scriptstyle{{\rm trg}~\nu} & & & & \scriptstyle{{\rm src}~\mu} \\
& &
\ar[dl] _{p_1}
\ar[dr] ^{p_0}
& &
\\
&
\ar[uul] ^{{\rm Leg1}({\rm Cod}~\nu)}
\ar[dr] _<>(0.4){{\rm Leg1}({\rm Cod}~\nu)\hspace{20pt}}
& &
\ar[dl] ^<>(0.4){\hspace{20pt}{\rm Leg1}({\rm Cod}~\mu)}
\ar[uur] _{{\rm Leg0}({\rm Cod}~\mu)}
\\
& & \scriptstyle{{\rm src}({\rm Cod}~\nu)} \;=\; {{\rm trg}({\rm Cod}~\mu)} & &
}$$
\<close>
definition hcomp
where "hcomp \<nu> \<mu> \<equiv> if arr \<mu> \<and> arr \<nu> \<and> src \<nu> = trg \<mu> then
\<lparr>Chn = chine_hcomp \<nu> \<mu>,
Dom = composite_span.this C prj0 prj1 (Dom \<nu>) (Dom \<mu>),
Cod = composite_span.this C prj0 prj1 (Cod \<nu>) (Cod \<mu>)\<rparr>
else
null"
notation hcomp (infixr "\<star>" 53)
lemma chine_hcomp_props:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "\<guillemotleft>chine_hcomp \<nu> \<mu> : Leg0 (Dom \<nu>) \<down>\<down> Leg1 (Dom \<mu>) \<rightarrow>\<^sub>C Leg0 (Cod \<nu>) \<down>\<down> Leg1 (Cod \<mu>)\<guillemotright>"
and "C.commutative_square (Leg0 (Cod \<nu>)) (Leg1 (Cod \<mu>))
(Chn \<nu> \<cdot> \<p>\<^sub>1[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)])
(Chn \<mu> \<cdot> \<p>\<^sub>0[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)])"
and "C.commutative_square \<p>\<^sub>1[Leg0 (Cod \<nu>), Leg1 (Cod \<mu>)] (Chn \<nu>)
(chine_hcomp \<nu> \<mu>) \<p>\<^sub>1[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)]"
and "C.commutative_square \<p>\<^sub>0[Leg0 (Cod \<nu>), Leg1 (Cod \<mu>)] (Chn \<mu>)
(chine_hcomp \<nu> \<mu>) \<p>\<^sub>0[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)]"
and "\<p>\<^sub>0[Leg0 (Cod \<nu>), Leg1 (Cod \<mu>)] \<cdot> chine_hcomp \<nu> \<mu> =
Chn \<mu> \<cdot> \<p>\<^sub>0[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)]"
and "\<p>\<^sub>1[Leg0 (Cod \<nu>), Leg1 (Cod \<mu>)] \<cdot> chine_hcomp \<nu> \<mu> =
Chn \<nu> \<cdot> \<p>\<^sub>1[Leg0 (Dom \<nu>), Leg1 (Dom \<mu>)]"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using assms arr_char by auto
show 0: "C.commutative_square \<nu>.cod.leg0 \<mu>.cod.leg1
(\<nu>.chine \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]) (\<mu>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1])"
using assms src_def trg_def C.pullback_commutes C.comp_reduce C.commutative_square_def
by auto
show 1: "\<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1] \<cdot> chine_hcomp \<nu> \<mu> = \<nu>.chine \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]"
unfolding chine_hcomp_def
using 0 by simp
show 2: "\<p>\<^sub>0[\<nu>.cod.leg0, \<mu>.cod.leg1] \<cdot> chine_hcomp \<nu> \<mu> = \<mu>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1]"
unfolding chine_hcomp_def
using 0 by simp
show 3: "\<guillemotleft>chine_hcomp \<nu> \<mu> : \<nu>.dom.leg0 \<down>\<down> \<mu>.dom.leg1 \<rightarrow>\<^sub>C \<nu>.cod.leg0 \<down>\<down> \<mu>.cod.leg1\<guillemotright>"
unfolding chine_hcomp_def
using assms 0 src_def trg_def C.tuple_in_hom by auto
show "C.commutative_square \<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1] \<nu>.chine
(chine_hcomp \<nu> \<mu>) \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]"
using assms src_def trg_def 1 3 by auto
show "C.commutative_square \<p>\<^sub>0[\<nu>.cod.leg0, \<mu>.cod.leg1] \<mu>.chine
(chine_hcomp \<nu> \<mu>) \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1]"
using assms src_def trg_def 2 3 by auto
qed
lemma chine_hcomp_in_hom [intro]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "\<guillemotleft>chine_hcomp \<nu> \<mu> : Leg0 (Dom \<nu>) \<down>\<down> Leg1 (Dom \<mu>) \<rightarrow>\<^sub>C Leg0 (Cod \<nu>) \<down>\<down> Leg1 (Cod \<mu>)\<guillemotright>"
using assms chine_hcomp_props(1) by simp
lemma arrow_of_spans_hcomp:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "arrow_of_spans C (\<nu> \<star> \<mu>)"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using assms arr_char by auto
show ?thesis
proof
show span_Dom: "C.span (Leg0 (Dom (\<nu> \<star> \<mu>))) (Leg1 (Dom (\<nu> \<star> \<mu>)))"
using assms src_def trg_def hcomp_def C.seqI' by auto
interpret Dom: span_in_category C \<open>Dom (\<nu> \<star> \<mu>)\<close>
using span_Dom by (unfold_locales, auto)
show span_Cod: "C.span (Leg0 (Cod (\<nu> \<star> \<mu>))) (Leg1 (Cod (\<nu> \<star> \<mu>)))"
using assms hcomp_def src_def trg_def by auto
interpret Cod: span_in_category C \<open>Cod (\<nu> \<star> \<mu>)\<close>
using span_Cod by (unfold_locales, auto)
show map: "\<guillemotleft>Chn (\<nu> \<star> \<mu>) : Dom.apex \<rightarrow>\<^sub>C Cod.apex\<guillemotright>"
using assms src_def trg_def chine_hcomp_props hcomp_def Cod.apex_def Dom.apex_def
by auto
show "Cod.leg0 \<cdot> Chn (\<nu> \<star> \<mu>) = Dom.leg0"
by (metis C.comp_assoc \<mu>.leg0_commutes arrow_of_spans_data.select_convs(1-3)
assms(1-3) chine_hcomp_props(5) hcomp_def span_data.select_convs(1))
show "Cod.leg1 \<cdot> Chn (\<nu> \<star> \<mu>) = Dom.leg1"
proof -
have "(\<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1]) \<cdot> chine_hcomp \<nu> \<mu> =
\<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]"
by (metis C.comp_assoc \<nu>.leg1_commutes assms(1-3) chine_hcomp_props(6))
thus ?thesis
using assms src_def trg_def hcomp_def chine_hcomp_props \<nu>.chine_in_hom C.comp_reduce
by auto
qed
qed
qed
lemma chine_hcomp_ide_arr:
assumes "ide f" and "arr \<mu>" and "src f = trg \<mu>"
shows "chine_hcomp f \<mu> =
\<langle>\<p>\<^sub>1[Leg0 (Dom f), Leg1 (Dom \<mu>)]
\<lbrakk>Leg0 (Cod f), Leg1 (Cod \<mu>)\<rbrakk>
Chn \<mu> \<cdot> \<p>\<^sub>0[Leg0 (Dom f), Leg1 (Dom \<mu>)]\<rangle>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret f: arrow_of_spans C f
using assms ide_char by auto
have 1: "C.cospan f.dom.leg0 \<mu>.dom.leg1"
using assms src_def trg_def by auto
have "chine_hcomp f \<mu> = \<langle>f.chine \<cdot> \<p>\<^sub>1[f.dom.leg0, \<mu>.dom.leg1]
\<lbrakk>f.cod.leg0, \<mu>.cod.leg1\<rbrakk>
\<mu>.chine \<cdot> \<p>\<^sub>0[f.dom.leg0, \<mu>.dom.leg1]\<rangle>"
unfolding chine_hcomp_def
using assms ide_char by simp
moreover have "f.chine \<cdot> \<p>\<^sub>1[f.dom.leg0, \<mu>.dom.leg1] = \<p>\<^sub>1[f.dom.leg0, \<mu>.dom.leg1]"
using assms 1 C.comp_ide_arr ide_char by auto
ultimately show ?thesis by argo
qed
lemma chine_hcomp_arr_ide:
assumes "arr \<mu>" and "ide f" and "src \<mu> = trg f"
shows "chine_hcomp \<mu> f =
\<langle>Chn \<mu> \<cdot> \<p>\<^sub>1[Leg0 (Dom \<mu>), Leg1 (Dom f)]
\<lbrakk>Leg0 (Cod \<mu>), Leg1 (Cod f)\<rbrakk>
\<p>\<^sub>0[Leg0 (Dom \<mu>), Leg1 (Dom f)]\<rangle>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret f: arrow_of_spans C f
using assms ide_char by auto
have 1: "C.cospan \<mu>.dom.leg0 f.dom.leg1"
using assms src_def trg_def by auto
have "chine_hcomp \<mu> f = \<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, f.dom.leg1]
\<lbrakk>\<mu>.cod.leg0, f.cod.leg1\<rbrakk>
f.chine \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, f.dom.leg1]\<rangle>"
unfolding chine_hcomp_def
using assms ide_char by simp
moreover have "f.chine \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, f.dom.leg1] = \<p>\<^sub>0[\<mu>.dom.leg0, f.dom.leg1]"
using assms 1 C.comp_ide_arr ide_char by auto
ultimately show ?thesis by argo
qed
lemma chine_hcomp_ide_ide:
assumes "ide g" and "ide f" and "src g = trg f"
shows "chine_hcomp g f = Leg0 (Dom g) \<down>\<down> Leg1 (Dom f)"
proof -
interpret g: identity_arrow_of_spans C g
using assms ide_char' by auto
interpret f: identity_arrow_of_spans C f
using assms ide_char' by auto
have 1: "C.cospan g.dom.leg0 f.dom.leg1"
using assms src_def trg_def by auto
have "chine_hcomp g f = \<langle>g.chine \<cdot> \<p>\<^sub>1[g.dom.leg0, f.dom.leg1]
\<lbrakk>g.cod.leg0, f.cod.leg1\<rbrakk>
\<p>\<^sub>0[g.dom.leg0, f.dom.leg1]\<rangle>"
using assms chine_hcomp_arr_ide by simp
moreover have "g.chine \<cdot> \<p>\<^sub>1[g.dom.leg0, f.dom.leg1] = \<p>\<^sub>1[g.dom.leg0, f.dom.leg1]"
using assms 1 C.comp_ide_arr ide_char by auto
ultimately have "chine_hcomp g f = \<langle>\<p>\<^sub>1[g.dom.leg0, f.dom.leg1]
\<lbrakk>g.cod.leg0, f.cod.leg1\<rbrakk>
\<p>\<^sub>0[g.dom.leg0, f.dom.leg1]\<rangle>"
by simp
also have "... =
\<langle>\<p>\<^sub>1[g.dom.leg0, f.dom.leg1] \<cdot> (g.dom.leg0 \<down>\<down> f.dom.leg1)
\<lbrakk>g.cod.leg0, f.cod.leg1\<rbrakk>
\<p>\<^sub>0[g.dom.leg0, f.dom.leg1] \<cdot> (g.dom.leg0 \<down>\<down> f.dom.leg1)\<rangle>"
using assms 1 C.comp_arr_dom by simp
also have "... = g.dom.leg0 \<down>\<down> f.dom.leg1"
using 1 C.pullback_commutes C.tuple_prj by simp
finally show ?thesis by simp
qed
lemma chine_hcomp_trg_arr:
assumes "arr \<mu>"
shows "chine_hcomp (trg \<mu>) \<mu> =
\<langle>\<p>\<^sub>1[C.cod (Leg1 (Dom \<mu>)), Leg1 (Dom \<mu>)]
\<lbrakk>C.cod (Leg1 (Dom \<mu>)), Leg1 (Cod \<mu>)\<rbrakk>
Chn \<mu> \<cdot> \<p>\<^sub>0[C.cod (Leg1 (Dom \<mu>)), Leg1 (Dom \<mu>)]\<rangle>"
using assms chine_hcomp_ide_arr ide_trg src_trg trg_def by force
lemma chine_hcomp_trg_ide:
assumes "ide f"
shows "chine_hcomp (trg f) f = C.cod (Leg1 (Dom f)) \<down>\<down> Leg1 (Dom f)"
using assms chine_hcomp_ide_ide ide_trg src_trg trg_def by force
lemma chine_hcomp_arr_src:
assumes "arr \<mu>"
shows "chine_hcomp \<mu> (src \<mu>) =
\<langle>Chn \<mu> \<cdot> \<p>\<^sub>1[Leg0 (Dom \<mu>), C.cod (Leg0 (Dom \<mu>))]
\<lbrakk>Leg0 (Cod \<mu>), C.cod (Leg0 (Dom \<mu>))\<rbrakk>
\<p>\<^sub>0[Leg0 (Dom \<mu>), C.cod (Leg0 (Dom \<mu>))]\<rangle>"
using assms chine_hcomp_arr_ide ide_src src_def trg_src by force
lemma chine_hcomp_ide_src:
assumes "ide f"
shows "chine_hcomp f (src f) = Leg0 (Dom f) \<down>\<down> C.cod (Leg0 (Dom f))"
using assms chine_hcomp_ide_ide src.preserves_ide src_def trg_src by force
lemma src_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "src (\<nu> \<star> \<mu>) = src \<mu>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using assms arr_char by auto
have "C.cod (\<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1]) = C.cod \<mu>.dom.leg0"
using assms C.commutative_squareE chine_hcomp_props(2)
by (metis (mono_tags, lifting) C.cod_comp C.match_3 \<mu>.leg0_commutes \<mu>.dom.is_span)
thus ?thesis
using assms arr_char hcomp_def src_def C.comp_cod_arr C.comp_arr_dom arrow_of_spans_hcomp
by simp
qed
lemma trg_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "trg (\<nu> \<star> \<mu>) = trg \<nu>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using assms arr_char by auto
have "C.cod (\<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]) = \<nu>.dtrg"
using assms C.commutative_squareE chine_hcomp_props(2)
by (metis (mono_tags, lifting) C.cod_comp C.match_3 \<nu>.leg1_commutes \<nu>.dom.is_span)
thus ?thesis
using assms arr_char hcomp_def trg_def C.comp_cod_arr C.comp_arr_dom arrow_of_spans_hcomp
by simp
qed
lemma dom_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "dom (\<nu> \<star> \<mu>) = dom \<nu> \<star> dom \<mu>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using assms arr_char by auto
interpret \<nu>\<mu>: arrow_of_spans C \<open>hcomp \<nu> \<mu>\<close>
using assms arr_char arrow_of_spans_hcomp by simp
have 1: "C.cospan \<mu>.dom.leg1 \<nu>.dom.leg0"
using assms \<mu>.dom.is_span \<nu>.dom.is_span src_def trg_def by auto
have "dom (\<nu> \<star> \<mu>) =
\<lparr>Chn = \<nu>.dom.leg0 \<down>\<down> \<mu>.dom.leg1,
Dom = \<lparr>Leg0 = \<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rparr>\<rparr>"
using assms \<nu>\<mu>.arrow_of_spans_axioms \<nu>\<mu>.dom.leg_simps(2) \<nu>\<mu>.dom.is_span
arr_char dom_char hcomp_def
by auto
also have "... =
\<lparr>Chn = chine_hcomp (dom \<nu>) (dom \<mu>),
Dom = \<lparr>Leg0 = \<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rparr>\<rparr>"
using assms src_dom trg_dom ide_dom dom_char chine_hcomp_ide_ide
apply auto
by (metis (no_types, lifting) arrow_of_spans_data.select_convs(2))
also have "... = dom \<nu> \<star> dom \<mu>"
using assms src_dom trg_dom arr_dom dom_char hcomp_def
apply auto
by (metis (no_types, lifting))
finally show ?thesis by blast
qed
lemma cod_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
shows "cod (\<nu> \<star> \<mu>) = cod \<nu> \<star> cod \<mu>"
proof -
interpret \<mu>: arrow_of_spans C \<mu>
using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu>
using assms arr_char by auto
interpret \<nu>\<mu>: arrow_of_spans C \<open>hcomp \<nu> \<mu>\<close>
using assms arr_char arrow_of_spans_hcomp by simp
have 1: "C.cospan \<mu>.cod.leg1 \<nu>.cod.leg0"
using assms \<mu>.cod.is_span \<nu>.cod.is_span src_def trg_def by simp
have 2: "cod (\<nu> \<star> \<mu>) =
\<lparr>Chn = \<nu>.cod.leg0 \<down>\<down> \<mu>.cod.leg1,
Dom = \<lparr>Leg0 = \<mu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<mu>.cod.leg1],
Leg1 = \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<mu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<mu>.cod.leg1],
Leg1 = \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1]\<rparr>\<rparr>"
using assms \<nu>\<mu>.arrow_of_spans_axioms \<nu>\<mu>.cod.leg_simps(2) \<nu>\<mu>.cod.is_span
arr_char cod_char hcomp_def
by auto
also have "... =
\<lparr>Chn = chine_hcomp (cod \<nu>) (cod \<mu>),
Dom = \<lparr>Leg0 = \<mu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<mu>.cod.leg1],
Leg1 = \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<mu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<mu>.cod.leg1],
Leg1 = \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<mu>.cod.leg1]\<rparr>\<rparr>"
using assms src_cod trg_cod ide_cod cod_char chine_hcomp_ide_ide
apply auto
by (metis (no_types, lifting) arrow_of_spans_data.select_convs(2))
also have "... = cod \<nu> \<star> cod \<mu>"
using assms src_cod trg_cod arr_cod cod_char hcomp_def
apply auto
by (metis (no_types, lifting))
finally show ?thesis by simp
qed
lemma hcomp_vcomp:
assumes "arr \<mu>" and "arr \<nu>" and "src \<nu> = trg \<mu>"
and "arr \<mu>'" and "arr \<nu>'" and "src \<nu>' = trg \<mu>'"
and "seq \<mu>' \<mu>" and "seq \<nu>' \<nu>"
shows "(\<nu>' \<bullet> \<nu>) \<star> (\<mu>' \<bullet> \<mu>) = (\<nu>' \<star> \<mu>') \<bullet> (\<nu> \<star> \<mu>)"
proof -
interpret \<mu>: arrow_of_spans C \<mu> using assms arr_char by auto
interpret \<nu>: arrow_of_spans C \<nu> using assms arr_char by auto
interpret \<mu>': arrow_of_spans C \<mu>' using assms arr_char by auto
interpret \<nu>': arrow_of_spans C \<nu>' using assms arr_char by auto
interpret \<nu>\<mu>: arrow_of_spans C \<open>hcomp \<nu> \<mu>\<close>
using assms arr_char arrow_of_spans_hcomp by auto
interpret \<nu>'\<mu>': arrow_of_spans C \<open>hcomp \<nu>' \<mu>'\<close>
using assms arr_char arrow_of_spans_hcomp by auto
have 1: "Dom \<nu>' = Cod \<nu> \<and> Dom \<mu>' = Cod \<mu>"
using assms src_def trg_def seq_char by blast
have 2: "Dom (\<mu>' \<bullet> \<mu>) = Dom \<mu> \<and> Dom (\<nu>' \<bullet> \<nu>) = Dom \<nu> \<and>
Cod (\<mu>' \<bullet> \<mu>) = Cod \<mu>' \<and> Cod (\<nu>' \<bullet> \<nu>) = Cod \<nu>'"
using assms seq_char arr_char vcomp_def
by (metis arrow_of_spans_data.select_convs(2) arrow_of_spans_data.select_convs(3))
have 3: "chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>) =
\<langle>Chn (\<nu>' \<bullet> \<nu>) \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]
\<lbrakk>\<nu>'.cod.leg0, \<mu>'.cod.leg1\<rbrakk>
Chn (\<mu>' \<bullet> \<mu>) \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rangle>"
unfolding chine_hcomp_def using 2 by simp
have C1: "C.commutative_square \<nu>'.cod.leg0 \<mu>'.cod.leg1
(Chn \<nu>' \<cdot> \<p>\<^sub>1[\<nu>'.dom.leg0, \<mu>'.dom.leg1])
(Chn \<mu>' \<cdot> \<p>\<^sub>0[\<nu>'.dom.leg0, \<mu>'.dom.leg1])"
using assms 1 vcomp_def seq_char arr_char chine_hcomp_props(2) by blast
have C2: "C.commutative_square \<nu>'.cod.leg0 \<mu>'.cod.leg1
(Chn (\<nu>' \<bullet> \<nu>) \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1])
(Chn (\<mu>' \<bullet> \<mu>) \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1])"
by (metis "2" assms(3,6-8) chine_hcomp_props(2) src.as_nat_trans.preserves_comp_2
trg.as_nat_trans.preserves_comp_2 vseq_implies_hpar(2))
have "(\<nu>' \<bullet> \<nu>) \<star> (\<mu>' \<bullet> \<mu>) =
\<lparr>Chn = chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>),
Dom = \<lparr>Leg0 = \<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<mu>'.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>'.cod.leg0, \<mu>'.cod.leg1],
Leg1 = \<nu>'.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>'.cod.leg0, \<mu>'.cod.leg1]\<rparr>\<rparr>"
using "2" assms(3,6-8) hcomp_def src_vcomp trg_vcomp by presburger
moreover have "(\<nu>' \<star> \<mu>') \<bullet> (\<nu> \<star> \<mu>) =
\<lparr>Chn = chine_hcomp \<nu>' \<mu>' \<cdot> chine_hcomp \<nu> \<mu>,
Dom = \<lparr>Leg0 = \<mu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<mu>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<mu>.dom.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<mu>'.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>'.cod.leg0, \<mu>'.cod.leg1],
Leg1 = \<nu>'.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>'.cod.leg0, \<mu>'.cod.leg1]\<rparr>\<rparr>"
using "1" \<nu>'\<mu>'.arrow_of_spans_axioms \<nu>\<mu>.arrow_of_spans_axioms assms(1-6)
hcomp_def span_vertical_category.seq_char span_vertical_category_axioms vcomp_eq
by fastforce
moreover have "chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>) = chine_hcomp \<nu>' \<mu>' \<cdot> chine_hcomp \<nu> \<mu>"
proof -
have "C.cospan \<nu>'.cod.leg0 \<mu>'.cod.leg1"
using assms src_def trg_def by simp
moreover have "C.seq \<p>\<^sub>1[\<nu>'.cod.leg0, \<mu>'.cod.leg1] (chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>))"
using assms 2 C2 chine_hcomp_props [of "\<mu>' \<bullet> \<mu>" "\<nu>' \<bullet> \<nu>"] by auto
moreover have "C.seq \<p>\<^sub>1[\<nu>'.cod.leg0, \<mu>'.cod.leg1] (chine_hcomp \<nu>' \<mu>' \<cdot> chine_hcomp \<nu> \<mu>)"
using assms 1 chine_hcomp_props [of \<mu> \<nu>] chine_hcomp_props [of \<mu>' \<nu>'] by auto
moreover have "\<p>\<^sub>0[\<nu>'.cod.leg0, \<mu>'.cod.leg1] \<cdot> chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>) =
\<p>\<^sub>0[\<nu>'.cod.leg0, \<mu>'.cod.leg1] \<cdot> chine_hcomp \<nu>' \<mu>' \<cdot> chine_hcomp \<nu> \<mu>"
by (metis (no_types, lifting) "1" "2" C.comp_assoc Chn_vcomp assms(1-8)
chine_hcomp_props(5) src_vcomp trg_vcomp)
moreover have "\<p>\<^sub>1[\<nu>'.cod.leg0, \<mu>'.cod.leg1] \<cdot> chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>) =
\<p>\<^sub>1[\<nu>'.cod.leg0, \<mu>'.cod.leg1] \<cdot> chine_hcomp \<nu>' \<mu>' \<cdot> chine_hcomp \<nu> \<mu>"
by (metis (no_types, lifting) "1" "3" C.comp_assoc C.prj_tuple(2) C1 C2 Chn_vcomp
assms(1-3,8) chine_hcomp_def chine_hcomp_props(6))
ultimately show ?thesis
using C.prj_joint_monic
[of "\<nu>'.cod.leg0" "\<mu>'.cod.leg1"
"chine_hcomp (\<nu>' \<bullet> \<nu>) (\<mu>' \<bullet> \<mu>)" "chine_hcomp \<nu>' \<mu>' \<cdot> chine_hcomp \<nu> \<mu>"]
by simp
qed
ultimately show ?thesis by auto
qed
interpretation H: "functor" VV.comp vcomp \<open>\<lambda>\<nu>\<mu>. fst \<nu>\<mu> \<star> snd \<nu>\<mu>\<close>
proof
show "\<And>\<nu>\<mu>. \<not> VV.arr \<nu>\<mu> \<Longrightarrow> fst \<nu>\<mu> \<star> snd \<nu>\<mu> = null"
- using hcomp_def VV.arr_char null_char by auto
+ using hcomp_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C null_char by auto
show "\<And>\<nu>\<mu>. VV.arr \<nu>\<mu> \<Longrightarrow> arr (fst \<nu>\<mu> \<star> snd \<nu>\<mu>)"
- using arr_char arrow_of_spans_hcomp VV.arr_char by simp
+ using arr_char arrow_of_spans_hcomp VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show "\<And>\<nu>\<mu>. VV.arr \<nu>\<mu> \<Longrightarrow>
dom (fst \<nu>\<mu> \<star> snd \<nu>\<mu>) = fst (VV.dom \<nu>\<mu>) \<star> snd (VV.dom \<nu>\<mu>)"
- using VV.arr_char VV.dom_char dom_hcomp by auto
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.dom_char\<^sub>S\<^sub>b\<^sub>C dom_hcomp by auto
show "\<And>\<nu>\<mu>. VV.arr \<nu>\<mu> \<Longrightarrow> cod (fst \<nu>\<mu> \<star> snd \<nu>\<mu>) = fst (VV.cod \<nu>\<mu>) \<star> snd (VV.cod \<nu>\<mu>)"
- using VV.arr_char VV.cod_char cod_hcomp by auto
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C cod_hcomp by auto
show "\<And>\<nu>\<mu>' \<nu>\<mu>. VV.seq \<nu>\<mu>' \<nu>\<mu> \<Longrightarrow> fst (VV.comp \<nu>\<mu>' \<nu>\<mu>) \<star> snd (VV.comp \<nu>\<mu>' \<nu>\<mu>) =
(fst \<nu>\<mu>' \<star> snd \<nu>\<mu>') \<bullet> (fst \<nu>\<mu> \<star> snd \<nu>\<mu>)"
proof -
fix \<nu>\<mu>' \<nu>\<mu>
assume 1: "VV.seq \<nu>\<mu>' \<nu>\<mu>"
have "VV.comp \<nu>\<mu>' \<nu>\<mu> = (fst \<nu>\<mu>' \<bullet> fst \<nu>\<mu>, snd \<nu>\<mu>' \<bullet> snd \<nu>\<mu>)"
- by (metis (no_types, lifting) "1" VV.comp_simp VV.seq_char VxV.comp_char VxV.seqE)
+ by (metis (no_types, lifting) "1" VV.comp_simp VV.seq_char\<^sub>S\<^sub>b\<^sub>C VxV.comp_char VxV.seqE)
thus "fst (VV.comp \<nu>\<mu>' \<nu>\<mu>) \<star> snd (VV.comp \<nu>\<mu>' \<nu>\<mu>) =
(fst \<nu>\<mu>' \<star> snd \<nu>\<mu>') \<bullet> (fst \<nu>\<mu> \<star> snd \<nu>\<mu>)"
- using 1 hcomp_vcomp VV.seq_char VV.arr_char VV.comp_char
+ using 1 hcomp_vcomp VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.comp_char
by (metis (no_types, lifting) fst_conv snd_conv)
qed
qed
lemma hcomp_is_functor:
shows "functor VV.comp vcomp (\<lambda>\<nu>\<mu>. fst \<nu>\<mu> \<star> snd \<nu>\<mu>)"
..
lemma ide_hcomp:
assumes "ide f" and "ide g" and "src f = trg g"
shows "ide (f \<star> g)"
- using assms VV.ide_char VV.arr_char H.preserves_ide [of "(f, g)"] by auto
+ using assms VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C H.preserves_ide [of "(f, g)"] by auto
sublocale horizontal_composition vcomp hcomp src trg
- using src_hcomp trg_hcomp VV.arr_char not_arr_null hcomp_def null_char
+ using src_hcomp trg_hcomp VV.arr_char\<^sub>S\<^sub>b\<^sub>C not_arr_null hcomp_def null_char
by (unfold_locales, auto)
lemma has_horizontal_composition:
shows "horizontal_composition vcomp hcomp src trg"
..
end
subsection "The Bicategory Span(C)"
context span_bicategory
begin
lemma arr_eqI:
assumes "par \<mu> \<mu>'" and "Chn \<mu> = Chn \<mu>'"
shows "\<mu> = \<mu>'"
using assms dom_char cod_char by auto
abbreviation \<l>
where "\<l> f \<equiv> \<lparr>Chn = \<p>\<^sub>0[C.cod (Leg1 (Dom f)), Leg1 (Dom f)],
Dom = Dom (L f), Cod = Cod f\<rparr>"
interpretation \<ll>: transformation_by_components vcomp vcomp L map \<l>
proof
have *: "\<And>f. ide f \<Longrightarrow> arrow_of_spans C (\<l> f)"
proof -
fix f
assume f: "ide f"
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
interpret \<l>f: arrow_of_spans C \<open>\<l> f\<close>
proof
show Dom: "C.span (Leg0 (Dom (\<l> f))) (Leg1 (Dom (\<l> f)))"
using f
by (simp add: arrow_of_spans_hcomp arrow_of_spans.axioms(2)
span_in_category.is_span)
interpret Dom: span_in_category C \<open>Dom (\<l> f)\<close>
using Dom by (unfold_locales, auto)
show Cod: "C.span (Leg0 (Cod (\<l> f))) (Leg1 (Cod (\<l> f)))"
using f hcomp_def trg_def src_def ide_mkObj C.pullback_commutes by force
interpret Cod: span_in_category C \<open>Cod (\<l> f)\<close>
using Cod by (unfold_locales, auto)
show "\<guillemotleft>Chn (\<l> f) : Dom.apex \<rightarrow>\<^sub>C Cod.apex\<guillemotright>"
proof -
have "C.dom Dom.leg0 = C.cod f.dom.leg1 \<down>\<down> f.dom.leg1"
proof -
have "arr (trg f)"
using f by simp
hence "Dom (\<l> f) = \<lparr>Leg0 = f.dom.leg0 \<cdot> \<p>\<^sub>0[C.cod f.dom.leg1, f.dom.leg1],
Leg1 = C.cod f.dom.leg1 \<cdot> \<p>\<^sub>1[C.cod f.dom.leg1, f.dom.leg1]\<rparr>"
using f src_def trg_def hcomp_def by simp
thus ?thesis
using f Dom hcomp_def by auto
qed
thus ?thesis
using f ide_char Dom.apex_def Cod.apex_def by simp
qed
show "Cod.leg0 \<cdot> Chn (\<l> f) = Dom.leg0"
using f ide_char hcomp_def src_def trg_def C.comp_arr_ide ide_mkObj by simp
show "Cod.leg1 \<cdot> Chn (\<l> f) = Dom.leg1"
using f ide_char hcomp_def src_def trg_def C.pullback_commutes ide_mkObj
C.comp_arr_ide
by (simp add: C.commutative_square_def)
qed
show "arrow_of_spans C (\<l> f)" ..
qed
show 0: "\<And>f. ide f \<Longrightarrow> \<guillemotleft>\<l> f : L f \<Rightarrow> map f\<guillemotright>"
proof -
fix f
assume f: "ide f"
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
interpret \<l>f: arrow_of_spans C \<open>\<l> f\<close>
using f * by blast
show "\<guillemotleft>\<l> f : L f \<Rightarrow> map f\<guillemotright>"
proof
show 1: "arr (\<l> f)"
using f * arr_char by blast
show "dom (\<l> f) = L f"
using f 1 dom_char ideD(2) by auto
show "cod (\<l> f) = map f"
using f 1 cod_char ideD(3) by auto
qed
qed
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> \<l> (cod \<mu>) \<bullet> L \<mu> = map \<mu> \<bullet> \<l> (dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu> arr_char by auto
interpret \<l>_dom_\<mu>: arrow_of_spans C \<open>\<l> (dom \<mu>)\<close>
using \<mu> * [of "dom \<mu>"] by fastforce
interpret \<l>_cod_\<mu>: arrow_of_spans C \<open>\<l> (cod \<mu>)\<close>
using \<mu> * [of "cod \<mu>"] by fastforce
interpret L\<mu>: arrow_of_spans C \<open>L \<mu>\<close>
using \<mu> arr_char by blast
show "\<l> (cod \<mu>) \<bullet> L \<mu> = map \<mu> \<bullet> \<l> (dom \<mu>)"
proof (intro arr_eqI)
show par: "par (\<l> (cod \<mu>) \<bullet> L \<mu>) (map \<mu> \<bullet> \<l> (dom \<mu>))"
using \<mu> 0 [of "dom \<mu>"] 0 [of "cod \<mu>"] by fastforce
show "Chn (\<l> (cod \<mu>) \<bullet> L \<mu>) = Chn (map \<mu> \<bullet> \<l> (dom \<mu>))"
proof -
have "Chn (\<l> (cod \<mu>) \<bullet> L \<mu>) =
\<p>\<^sub>0[\<mu>.dtrg, \<mu>.cod.leg1] \<cdot>
\<langle>\<p>\<^sub>1[\<mu>.dtrg, \<mu>.dom.leg1] \<lbrakk>\<mu>.dtrg, \<mu>.cod.leg1\<rbrakk> \<mu>.chine \<cdot> \<p>\<^sub>0[\<mu>.dtrg, \<mu>.dom.leg1]\<rangle>"
proof -
have "Chn (\<l> (cod \<mu>) \<bullet> L \<mu>) = \<p>\<^sub>0[\<mu>.dtrg, \<mu>.cod.leg1] \<cdot> Chn (trg \<mu> \<star> \<mu>)"
using Chn_vcomp \<mu> cod_char par by force
moreover
have "Chn (trg \<mu> \<star> \<mu>) = \<langle>\<p>\<^sub>1[\<mu>.dtrg, \<mu>.dom.leg1]
\<lbrakk>\<mu>.dtrg, \<mu>.cod.leg1\<rbrakk>
\<mu>.chine \<cdot> \<p>\<^sub>0[\<mu>.dtrg, \<mu>.dom.leg1]\<rangle>"
using \<mu> hcomp_def chine_hcomp_trg_arr by simp
ultimately show ?thesis
using \<mu> by (auto simp add: cod_char)
qed
also have "... = \<mu>.chine \<cdot> \<p>\<^sub>0[C.cod \<mu>.dom.leg1, \<mu>.dom.leg1]"
using \<mu> C.in_homE C.pullback_commutes [of "C.cod \<mu>.dom.leg1" "\<mu>.dom.leg1"]
C.comp_reduce ide_char C.prj_tuple(1)
by auto
also have "... = Chn (map \<mu> \<bullet> \<l> (dom \<mu>))"
using \<mu> par seq_char dom_char vcomp_eq map_simp by simp
finally show ?thesis by blast
qed
qed
qed
qed
interpretation \<ll>: natural_isomorphism vcomp vcomp L map \<ll>.map
proof
fix f
assume f: "ide f"
show "iso (\<ll>.map f)"
proof -
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
have 1: "\<ll>.map f = \<lparr>Chn = \<p>\<^sub>0[f.dtrg, f.leg1], Dom = Dom (trg f \<star> f), Cod = Dom f\<rparr>"
using f ide_char cod_char by simp
interpret \<l>f: arrow_of_spans C \<open>\<ll>.map f\<close>
using f arr_char \<ll>.preserves_reflects_arr by fastforce
let ?\<l>f' = "\<lparr>Chn = \<langle>f.leg1 \<lbrakk>f.dtrg, f.leg1\<rbrakk> C.dom f.leg1\<rangle>,
Dom = Dom f, Cod = Dom (trg f \<star> f)\<rparr>"
have 2: "C.inverse_arrows \<l>f.chine (Chn ?\<l>f')"
using 1 C.pullback_arr_cod(2) [of "f.leg1"] by simp
interpret \<l>f': arrow_of_spans C ?\<l>f'
proof
show Dom: "C.span (Leg0 (Dom ?\<l>f')) (Leg1 (Dom ?\<l>f'))"
using f 1 by auto
interpret Dom: span_in_category C \<open>Dom ?\<l>f'\<close>
using Dom by (unfold_locales, auto)
show Cod: "C.span (Leg0 (Cod ?\<l>f')) (Leg1 (Cod ?\<l>f'))"
using f 1 \<l>f.dom.is_span by auto
interpret Cod: span_in_category C \<open>Cod ?\<l>f'\<close>
using Cod by (unfold_locales, auto)
show "\<guillemotleft>Chn ?\<l>f' : Dom.apex \<rightarrow>\<^sub>C Cod.apex\<guillemotright>"
using f src_def trg_def hcomp_def ide_mkObj Cod.apex_def Dom.apex_def
C.comp_arr_dom C.comp_cod_arr
by auto
show "Cod.leg0 \<cdot> Chn ?\<l>f' = Dom.leg0"
using 2 \<l>f.leg0_commutes C.invert_side_of_triangle
by (metis (no_types, lifting) "1" C.inverse_unique C.isoI \<l>f.dom.is_span
arrow_of_spans_data.select_convs(2) arrow_of_spans_data.select_convs(3))
show "Cod.leg1 \<cdot> Chn ?\<l>f' = Dom.leg1"
using 2 \<l>f.leg1_commutes C.invert_side_of_triangle
by (metis (no_types, lifting) "1" C.inverse_unique C.isoI \<l>f.dom.is_span
arrow_of_spans_data.select_convs(2) arrow_of_spans_data.select_convs(3))
qed
have "inverse_arrows (\<ll>.map f) ?\<l>f'"
proof
show "ide (?\<l>f' \<bullet> \<ll>.map f)"
proof -
have "?\<l>f' \<bullet> \<ll>.map f = dom (\<ll>.map f)"
proof -
have "?\<l>f' \<bullet> \<ll>.map f =
\<lparr>Chn = f.dtrg \<down>\<down> f.leg1, Dom = Dom (\<ll>.map f), Cod = Dom (\<ll>.map f)\<rparr>"
using f 1 2 f.arrow_of_spans_axioms \<l>f.arrow_of_spans_axioms
\<l>f'.arrow_of_spans_axioms vcomp_def ide_char arr_char
by (simp add: vcomp_def C.comp_inv_arr)
also have "... = dom (\<ll>.map f)"
by (metis (no_types, lifting) "1" C.pbdom_def \<l>f.chine_simps(2)
\<ll>.preserves_reflects_arr arrow_of_spans_data.select_convs(1)
dom_char f ideD(1))
finally show ?thesis by blast
qed
thus ?thesis
using \<l>f.arrow_of_spans_axioms arr_char by simp
qed
show "ide (\<ll>.map f \<bullet> ?\<l>f')"
proof -
have "\<ll>.map f \<bullet> ?\<l>f' = dom ?\<l>f'"
proof -
have "\<ll>.map f \<bullet> ?\<l>f' = \<lparr>Chn = Chn f, Dom = Dom ?\<l>f', Cod = Dom ?\<l>f'\<rparr>"
using f 1 2 f.arrow_of_spans_axioms \<l>f.arrow_of_spans_axioms
\<l>f'.arrow_of_spans_axioms vcomp_def ide_char arr_char
by fastforce
also have "... = dom ?\<l>f'"
using 1 \<l>f'.arrow_of_spans_axioms arr_char dom_char by simp
finally show ?thesis by blast
qed
thus ?thesis
using \<l>f'.arrow_of_spans_axioms arr_char by simp
qed
qed
thus ?thesis by auto
qed
qed
lemma \<ll>_is_natural_isomorphism:
shows "natural_isomorphism vcomp vcomp L map \<ll>.map"
..
sublocale L: equivalence_functor vcomp vcomp L
using L.isomorphic_to_identity_is_equivalence \<ll>.natural_isomorphism_axioms by simp
lemma equivalence_functor_L:
shows "equivalence_functor vcomp vcomp L"
..
abbreviation \<r>
where "\<r> f \<equiv> \<lparr>Chn = \<p>\<^sub>1[Leg0 (Dom f), C.cod (Leg0 (Dom f))],
Dom = Dom (R f), Cod = Cod f\<rparr>"
interpretation \<rho>: transformation_by_components vcomp vcomp R map \<r>
proof
have *: "\<And>f. ide f \<Longrightarrow> arrow_of_spans C (\<r> f)"
proof -
fix f
assume f: "ide f"
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
interpret \<r>f: arrow_of_spans C \<open>\<r> f\<close>
proof
show Dom: "C.span (Leg0 (Dom (\<r> f))) (Leg1 (Dom (\<r> f)))"
using f
by (simp add: arrow_of_spans_hcomp arrow_of_spans.axioms(2)
span_in_category.is_span)
interpret Dom: span_in_category C \<open>Dom (\<r> f)\<close>
using Dom by (unfold_locales, auto)
show Cod: "C.span (Leg0 (Cod (\<r> f))) (Leg1 (Cod (\<r> f)))"
using f hcomp_def trg_def src_def ide_mkObj C.pullback_commutes by force
interpret Cod: span_in_category C \<open>Cod (\<r> f)\<close>
using Cod by (unfold_locales, auto)
show "\<guillemotleft>Chn (\<r> f) : Dom.apex \<rightarrow>\<^sub>C Cod.apex\<guillemotright>"
proof -
have "C.dom Dom.leg0 = f.dom.leg0 \<down>\<down> C.cod f.dom.leg0"
proof -
have "arr (src f)"
using f by simp
hence "Dom (\<r> f) = \<lparr>Leg0 = C.cod f.dom.leg0 \<cdot> \<p>\<^sub>0[f.dom.leg0, C.cod f.dom.leg0],
Leg1 = f.dom.leg1 \<cdot> \<p>\<^sub>1[f.dom.leg0, C.cod f.dom.leg0]\<rparr>"
using f src_def trg_def by (simp add: hcomp_def)
thus ?thesis
using f ide_char Dom.apex_def Cod.apex_def by simp
qed
thus ?thesis
using f ide_char Dom.apex_def Cod.apex_def by simp
qed
show "Cod.leg0 \<cdot> Chn (\<r> f) = Dom.leg0"
using f ide_char hcomp_def src_def trg_def C.pullback_commutes
ide_mkObj C.comp_arr_ide
by (simp add: C.commutative_square_def)
show "Cod.leg1 \<cdot> Chn (\<r> f) = Dom.leg1"
using f ide_char hcomp_def src_def trg_def ide_mkObj C.comp_arr_ide
by (simp add: C.commutative_square_def)
qed
show "arrow_of_spans C (\<r> f)" ..
qed
show 0: "\<And>f. ide f \<Longrightarrow> \<guillemotleft>\<r> f : R f \<Rightarrow> map f\<guillemotright>"
proof -
fix f
assume f: "ide f"
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
interpret \<r>f: arrow_of_spans C \<open>\<r> f\<close>
using f * by blast
show "\<guillemotleft>\<r> f : R f \<Rightarrow> map f\<guillemotright>"
proof
show 1: "arr (\<r> f)"
using f * arr_char by blast
show "dom (\<r> f) = R f"
using f 1 dom_char ideD(2) by auto
show "cod (\<r> f) = map f"
using f 1 cod_char ideD(3) by auto
qed
qed
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> \<r> (cod \<mu>) \<bullet> R \<mu> = map \<mu> \<bullet> \<r> (dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
interpret \<mu>: arrow_of_spans C \<mu>
using \<mu> arr_char by auto
interpret \<r>_dom_\<mu>: arrow_of_spans C \<open>\<r> (dom \<mu>)\<close>
using \<mu> * [of "dom \<mu>"] by fastforce
interpret \<r>_cod_\<mu>: arrow_of_spans C \<open>\<r> (cod \<mu>)\<close>
using \<mu> * [of "cod \<mu>"] by fastforce
interpret R\<mu>: arrow_of_spans C \<open>R \<mu>\<close>
using \<mu> arr_char by blast
show "\<r> (cod \<mu>) \<bullet> R \<mu> = map \<mu> \<bullet> \<r> (dom \<mu>)"
proof (intro arr_eqI)
show par: "par (\<r> (cod \<mu>) \<bullet> R \<mu>) (map \<mu> \<bullet> \<r> (dom \<mu>))"
using \<mu> 0 [of "dom \<mu>"] 0 [of "cod \<mu>"] by force
show "Chn (\<r> (cod \<mu>) \<bullet> R \<mu>) = Chn (map \<mu> \<bullet> \<r> (dom \<mu>))"
proof -
have "Chn (\<r> (cod \<mu>) \<bullet> R \<mu>) =
\<p>\<^sub>1[\<mu>.cod.leg0, \<mu>.cod.src] \<cdot>
\<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<mu>.dsrc] \<lbrakk>\<mu>.cod.leg0, \<mu>.cod.src\<rbrakk> \<p>\<^sub>0[\<mu>.dom.leg0, \<mu>.dsrc]\<rangle>"
proof -
have "Chn (\<r> (cod \<mu>) \<bullet> R \<mu>) = \<p>\<^sub>1[\<mu>.cod.leg0, \<mu>.cod.src] \<cdot> Chn (\<mu> \<star> src \<mu>)"
using Chn_vcomp \<mu> cod_char par by force
moreover
have "Chn (\<mu> \<star> src \<mu>) = \<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<mu>.dsrc]
\<lbrakk>\<mu>.cod.leg0, \<mu>.dsrc\<rbrakk>
\<p>\<^sub>0[\<mu>.dom.leg0, \<mu>.dsrc]\<rangle>"
using \<mu> hcomp_def chine_hcomp_arr_src by simp
ultimately show ?thesis
using \<mu> by (auto simp add: cod_char)
qed
also have "... = \<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, C.cod \<mu>.dom.leg0]"
using \<mu> ide_char C.prj_tuple(2)
C.in_homE C.pullback_commutes [of "\<mu>.dom.leg0" "C.cod \<mu>.dom.leg0"]
C.comp_reduce
by auto
also have "... = Chn (map \<mu> \<bullet> \<r> (dom \<mu>))"
using \<mu> par seq_char dom_char vcomp_eq map_simp by simp
finally show ?thesis by blast
qed
qed
qed
qed
interpretation \<rho>: natural_isomorphism vcomp vcomp R map \<rho>.map
proof
fix f
assume f: "ide f"
show "iso (\<rho>.map f)"
proof -
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
have 1: "\<rho>.map f = \<lparr>Chn = \<p>\<^sub>1[f.leg0, f.dsrc], Dom = Dom (f \<star> src f), Cod = Dom f\<rparr>"
using f ide_char by auto
interpret \<rho>f: arrow_of_spans C \<open>\<rho>.map f\<close>
using f arr_char \<rho>.preserves_reflects_arr by fastforce
let ?\<rho>f' = "\<lparr>Chn = \<langle>C.dom f.leg0 \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>,
Dom = Dom f, Cod = Dom (f \<star> src f)\<rparr>"
have 2: "C.inverse_arrows (Chn (\<rho>.map f)) (Chn ?\<rho>f')"
using 1 C.pullback_arr_cod(1) [of "f.dom.leg0"] by simp
interpret \<rho>f': arrow_of_spans C ?\<rho>f'
proof
show Dom: "C.span (Leg0 (Dom ?\<rho>f')) (Leg1 (Dom ?\<rho>f'))"
using f 1 by auto
interpret Dom: span_in_category C \<open>Dom ?\<rho>f'\<close>
using Dom by (unfold_locales, auto)
show Cod: "C.span (Leg0 (Cod ?\<rho>f')) (Leg1 (Cod ?\<rho>f'))"
using f 1 \<rho>f.dom.is_span by auto
interpret Cod: span_in_category C \<open>Cod ?\<rho>f'\<close>
using Cod by (unfold_locales, auto)
show "\<guillemotleft>Chn ?\<rho>f' : Dom.apex \<rightarrow>\<^sub>C Cod.apex\<guillemotright>"
using f src_def trg_def hcomp_def ide_mkObj Cod.apex_def Dom.apex_def
C.comp_arr_dom C.comp_cod_arr
by auto
show "Cod.leg0 \<cdot> Chn ?\<rho>f' = Dom.leg0"
using 2 \<rho>f.leg0_commutes C.invert_side_of_triangle
by (metis (no_types, lifting) "1" C.inverse_unique C.isoI \<rho>f.dom.is_span
arrow_of_spans_data.select_convs(2) arrow_of_spans_data.select_convs(3))
show "Cod.leg1 \<cdot> Chn ?\<rho>f' = Dom.leg1"
using 2 \<rho>f.leg1_commutes C.invert_side_of_triangle
by (metis (no_types, lifting) "1" C.inverse_unique C.isoI \<rho>f.dom.is_span
arrow_of_spans_data.select_convs(2) arrow_of_spans_data.select_convs(3))
qed
have "inverse_arrows (\<rho>.map f) ?\<rho>f'"
proof
show "ide (?\<rho>f' \<bullet> \<rho>.map f)"
using "2" C.comp_inv_arr \<rho>f'.arrow_of_spans_axioms \<rho>f.arrow_of_spans_axioms
\<rho>f.chine_simps(2) dom_char f ideD(2) vcomp_def
by force
show "ide (\<rho>.map f \<bullet> ?\<rho>f')"
proof -
have "\<rho>.map f \<bullet> ?\<rho>f' = dom ?\<rho>f'"
proof -
have "\<rho>.map f \<bullet> ?\<rho>f' = \<lparr>Chn = Chn f, Dom = Dom ?\<rho>f', Cod = Dom ?\<rho>f'\<rparr>"
using f 1 2 f.arrow_of_spans_axioms
\<rho>f.arrow_of_spans_axioms \<rho>f'.arrow_of_spans_axioms
vcomp_def ide_char arr_char
by fastforce
also have "... = dom ?\<rho>f'"
using 1 \<rho>f'.arrow_of_spans_axioms arr_char dom_char by simp
finally show ?thesis by blast
qed
thus ?thesis
using \<rho>f'.arrow_of_spans_axioms arr_char by simp
qed
qed
thus ?thesis by auto
qed
qed
lemma \<rho>_is_natural_isomorphism:
shows "natural_isomorphism vcomp vcomp R map \<rho>.map"
..
sublocale R: equivalence_functor vcomp vcomp R
using R.isomorphic_to_identity_is_equivalence \<rho>.natural_isomorphism_axioms by simp
lemma equivalence_functor_R:
shows "equivalence_functor vcomp vcomp R"
..
definition unit ("\<i>[_]")
where "\<i>[a] \<equiv> \<lparr>Chn = \<p>\<^sub>0[Chn a, Chn a], Dom = Dom (a \<star> a), Cod = Cod a\<rparr>"
lemma unit_in_hom [intro]:
assumes "obj a"
shows "in_hhom \<i>[a] a a"
and "\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
proof -
show "\<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
proof (intro in_homI)
interpret a: identity_arrow_of_spans C a
using assms obj_char ide_char' by auto
have 0: "src a = trg a"
using assms arr_char obj_char src_def trg_def by (elim objE, auto)
interpret aa: arrow_of_spans C \<open>a \<star> a\<close>
using assms 0 a.arrow_of_spans_axioms arrow_of_spans_hcomp by auto
interpret aa: identity_arrow_of_spans C \<open>a \<star> a\<close>
proof
have "ide (a \<star> a)"
using assms 0 obj_char H.preserves_ide by simp
thus "C.ide aa.chine" using ide_char by auto
qed
have 1: "\<guillemotleft>\<p>\<^sub>0[a.chine, a.chine] : a.chine \<down>\<down> a.chine \<rightarrow>\<^sub>C a.chine\<guillemotright> \<and>
\<guillemotleft>\<p>\<^sub>1[a.chine, a.chine] : a.chine \<down>\<down> a.chine \<rightarrow>\<^sub>C a.chine\<guillemotright>"
by auto
have 2: "a.dom.leg0 = a.chine \<and> a.dom.leg1 = a.chine \<and>
a.cod.leg0 = a.chine \<and> a.cod.leg1 = a.chine"
using assms obj_char by (cases a, simp_all)
have 3: "a \<star> a = \<lparr>Chn = a.chine \<down>\<down> a.chine,
Dom = \<lparr>Leg0 = \<p>\<^sub>0[a.chine, a.chine], Leg1 = \<p>\<^sub>1[a.chine, a.chine]\<rparr>,
Cod = \<lparr>Leg0 = \<p>\<^sub>0[a.chine, a.chine], Leg1 = \<p>\<^sub>1[a.chine, a.chine]\<rparr>\<rparr>"
using assms 0 1 2 chine_hcomp_ide_ide hcomp_def C.comp_cod_arr
a.identity_arrow_of_spans_axioms ide_char'
by auto
have "aa.apex = a.chine \<down>\<down> a.chine"
using 3 aa.chine_eq_apex by auto
interpret \<i>a: arrow_of_spans C \<open>\<i>[a]\<close>
proof
have 4: "Dom \<i>[a] = Dom (a \<star> a)"
using assms hcomp_def unit_def by simp
have 5: "Cod \<i>[a] = Cod a"
using assms unit_def by simp
show Dom: "C.span (Leg0 (Dom \<i>[a])) (Leg1 (Dom \<i>[a]))"
using 4 by simp
interpret Dom: span_in_category C \<open>Dom \<i>[a]\<close>
using Dom by (unfold_locales, auto)
show Cod: "C.span (Leg0 (Cod \<i>[a])) (Leg1 (Cod \<i>[a]))"
using 5 by simp
interpret Cod: span_in_category C \<open>Cod \<i>[a]\<close>
using Cod by (unfold_locales, auto)
show "\<guillemotleft>Chn \<i>[a] : Dom.apex \<rightarrow>\<^sub>C Cod.apex\<guillemotright>"
by (simp add: \<open>aa.apex = a.chine \<down>\<down> a.chine\<close> unit_def)
show "Cod.leg0 \<cdot> Chn \<i>[a] = Dom.leg0"
unfolding unit_def using 1 2 3 C.comp_cod_arr by auto
show "Cod.leg1 \<cdot> Chn \<i>[a] = Dom.leg1"
unfolding unit_def using 1 2 3 C.comp_cod_arr C.pullback_ide_self by auto
qed
show "arr \<i>[a]"
using \<i>a.arrow_of_spans_axioms arr_char by simp
show "dom \<i>[a] = hcomp a a"
using 3 unit_def \<i>a.arrow_of_spans_axioms arr_char dom_char \<i>a.dom.apex_def
by auto
show "cod \<i>[a] = a"
using assms 3 obj_char arr_char dom_char cod_char unit_def
\<i>a.arrow_of_spans_axioms
by auto
qed
thus "in_hhom \<i>[a] a a"
using assms
by (metis arrI in_hhom_def objE vconn_implies_hpar(1) vconn_implies_hpar(2-4))
qed
lemma unit_simps [simp]:
assumes "obj a"
shows "src \<i>[a] = a" and "trg \<i>[a] = a"
and "dom \<i>[a] = hcomp a a" and "cod \<i>[a] = a"
using assms unit_in_hom by auto
lemma iso_unit:
assumes "obj a"
shows "iso \<i>[a]"
proof -
have "Chn \<i>[a] = \<p>\<^sub>0[Chn a, Chn a]"
unfolding unit_def by simp
moreover have "C.iso \<p>\<^sub>0[Chn a, Chn a]"
using assms C.ide_is_iso C.iso_is_arr C.iso_pullback_ide ide_char by blast
ultimately show ?thesis
using assms unit_in_hom iso_char by auto
qed
end
locale two_composable_arrows_of_spans =
span_bicategory +
\<mu>: arrow_of_spans C \<mu> +
\<nu>: arrow_of_spans C \<nu>
for \<mu> (structure)
and \<nu> (structure) +
assumes composable: "src \<mu> = trg \<nu>"
begin
lemma are_arrows [simp]:
shows "arr \<mu>" and "arr \<nu>"
using arr_char \<mu>.arrow_of_spans_axioms \<nu>.arrow_of_spans_axioms by auto
lemma legs_form_cospan:
shows "C.cospan \<mu>.dom.leg0 \<nu>.dom.leg1" and "C.cospan \<mu>.cod.leg0 \<nu>.cod.leg1"
using composable src_def trg_def by auto
interpretation \<mu>\<nu>: arrow_of_spans C \<open>\<mu> \<star> \<nu>\<close>
using arrow_of_spans_hcomp composable by auto
lemma composite_is_arrow [simp]:
shows "arr (\<mu> \<star> \<nu>)"
using \<mu>\<nu>.arrow_of_spans_axioms arr_char by auto
lemma composite_in_hom [intro]:
shows "\<guillemotleft>\<mu> \<star> \<nu> : dom \<mu> \<star> dom \<nu> \<Rightarrow> cod \<mu> \<star> cod \<nu>\<guillemotright>"
using composable by auto
lemma composite_simps [simp]:
shows "src (\<mu> \<star> \<nu>) = src \<nu>" and "trg (\<mu> \<star> \<nu>) = trg \<mu>"
and "dom (\<mu> \<star> \<nu>) = dom \<mu> \<star> dom \<nu>" and "cod (\<mu> \<star> \<nu>) = cod \<mu> \<star> cod \<nu>"
by (simp_all add: composable)
lemma chine_composite:
shows "Chn (\<mu> \<star> \<nu>) = \<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1]
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk>
\<nu>.chine \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1]\<rangle>"
unfolding hcomp_def chine_hcomp_def using composable by simp
lemma chine_composite_in_hom [intro]:
shows "\<guillemotleft>Chn (\<mu> \<star> \<nu>) : \<mu>.dom.leg0 \<down>\<down> \<nu>.dom.leg1 \<rightarrow>\<^sub>C \<mu>.cod.leg0 \<down>\<down> \<nu>.cod.leg1\<guillemotright>"
using hcomp_def chine_hcomp_props(1) composable by auto
end
sublocale two_composable_arrows_of_spans \<subseteq> arrow_of_spans C \<open>\<mu> \<star> \<nu>\<close>
proof -
interpret Dom\<mu>_Dom\<nu>: composite_span C prj0 prj1 \<open>Dom \<mu>\<close> \<open>Dom \<nu>\<close>
using legs_form_cospan(1) by (unfold_locales, auto)
interpret Cod\<mu>_Cod\<nu>: composite_span C prj0 prj1 \<open>Cod \<mu>\<close> \<open>Cod \<nu>\<close>
using legs_form_cospan(1) by (unfold_locales, auto)
interpret Dom_\<mu>\<nu>: span_in_category C \<open>Dom (\<mu> \<star> \<nu>)\<close>
apply unfold_locales apply (unfold hcomp_def)
using Dom\<mu>_Dom\<nu>.apex_def Dom\<mu>_Dom\<nu>.leg_simps(1) are_arrows(1) composable by auto
interpret Cod_\<mu>\<nu>: span_in_category C \<open>Cod (\<mu> \<star> \<nu>)\<close>
apply unfold_locales apply (unfold hcomp_def)
using Cod\<mu>_Cod\<nu>.apex_def Cod\<mu>_Cod\<nu>.leg_simps(1) are_arrows(1) composable by auto
show "arrow_of_spans C (\<mu> \<star> \<nu>)"
proof
show "\<guillemotleft>Chn (hcomp \<mu> \<nu>) : Dom_\<mu>\<nu>.apex \<rightarrow>\<^sub>C Cod_\<mu>\<nu>.apex\<guillemotright>"
unfolding hcomp_def
using are_arrows(1) are_arrows(2) arrow_of_spans_hcomp composable hcomp_def
arrow_of_spans.chine_in_hom Cod\<mu>_Cod\<nu>.leg_simps(4) Dom\<mu>_Dom\<nu>.leg_simps(3)
Dom\<mu>_Dom\<nu>.leg_simps(4) chine_composite_in_hom
by auto
show "Cod_\<mu>\<nu>.leg0 \<cdot> Chn (hcomp \<mu> \<nu>) = Dom_\<mu>\<nu>.leg0"
proof (unfold hcomp_def)
have "arrow_of_spans C
\<lparr>Chn = chine_hcomp \<mu> \<nu>, Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>"
using arrow_of_spans_hcomp composable hcomp_def by auto
thus "Leg0 (Cod (if arr \<nu> \<and> arr \<mu> \<and> src \<mu> = trg \<nu> then
\<lparr>Chn = chine_hcomp \<mu> \<nu>,
Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>
else null)) \<cdot>
Chn (if arr \<nu> \<and> arr \<mu> \<and> src \<mu> = trg \<nu> then
\<lparr>Chn = chine_hcomp \<mu> \<nu>,
Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>
else null) =
Leg0 (Dom (if arr \<nu> \<and> arr \<mu> \<and> src \<mu> = trg \<nu> then
\<lparr>Chn = chine_hcomp \<mu> \<nu>,
Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>
else null))"
using arrow_of_spans.leg0_commutes composable by fastforce
qed
show "Cod_\<mu>\<nu>.leg1 \<cdot> Chn (hcomp \<mu> \<nu>) = Dom_\<mu>\<nu>.leg1"
proof (unfold hcomp_def)
have "arrow_of_spans C
\<lparr>Chn = chine_hcomp \<mu> \<nu>, Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>"
using arrow_of_spans_hcomp composable hcomp_def by force
thus "Leg1 (Cod (if arr \<nu> \<and> arr \<mu> \<and> src \<mu> = trg \<nu> then
\<lparr>Chn = chine_hcomp \<mu> \<nu>,
Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>
else null)) \<cdot>
Chn (if arr \<nu> \<and> arr \<mu> \<and> src \<mu> = trg \<nu> then
\<lparr>Chn = chine_hcomp \<mu> \<nu>,
Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>
else null) =
Leg1 (Dom (if arr \<nu> \<and> arr \<mu> \<and> src \<mu> = trg \<nu> then
\<lparr>Chn = chine_hcomp \<mu> \<nu>,
Dom = Dom\<mu>_Dom\<nu>.this, Cod = Cod\<mu>_Cod\<nu>.this\<rparr>
else null))"
using arrow_of_spans.leg1_commutes composable by force
qed
qed
qed
locale two_composable_identity_arrows_of_spans =
two_composable_arrows_of_spans +
\<mu>: identity_arrow_of_spans C \<mu> +
\<nu>: identity_arrow_of_spans C \<nu>
begin
lemma are_identities [simp]:
shows "ide \<mu>" and "ide \<nu>"
using ide_char \<mu>.arrow_of_spans_axioms \<nu>.arrow_of_spans_axioms by auto
interpretation H: "functor" VV.comp vcomp \<open>\<lambda>\<nu>\<mu>. fst \<nu>\<mu> \<star> snd \<nu>\<mu>\<close>
using hcomp_is_functor by auto
interpretation \<mu>\<nu>: identity_arrow_of_spans C \<open>\<mu> \<star> \<nu>\<close>
using are_identities(1-2) composable ide_char' by blast
lemma ide_composite [simp]:
shows "ide (\<mu> \<star> \<nu>)"
using \<mu>\<nu>.identity_arrow_of_spans_axioms arrow_of_spans_axioms ide_char by auto
lemma apex_composite:
shows "\<mu>\<nu>.apex = \<mu>.leg0 \<down>\<down> \<nu>.leg1"
using dom.apex_def hcomp_def chine_hcomp_ide_ide composable legs_form_cospan
by simp
lemma leg0_composite:
shows "\<mu>\<nu>.leg0 = \<nu>.leg0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1]"
using dom.apex_def hcomp_def composable legs_form_cospan by simp
lemma leg1_composite:
shows "\<mu>\<nu>.leg1 = \<mu>.leg1 \<cdot> \<p>\<^sub>1[\<mu>.leg0, \<nu>.leg1]"
using dom.apex_def hcomp_def composable legs_form_cospan by simp
lemma chine_composite:
shows "Chn (\<mu> \<star> \<nu>) = \<mu>.leg0 \<down>\<down> \<nu>.leg1"
unfolding hcomp_def using chine_hcomp_ide_ide composable by simp
abbreviation prj\<^sub>0
where "prj\<^sub>0 \<equiv> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1]"
abbreviation prj\<^sub>1
where "prj\<^sub>1 \<equiv> \<p>\<^sub>1[\<mu>.leg0, \<nu>.leg1]"
lemma prj_in_hom [intro]:
shows "\<guillemotleft>prj\<^sub>1 : \<mu>.leg0 \<down>\<down> \<nu>.leg1 \<rightarrow>\<^sub>C \<mu>.apex\<guillemotright>"
and "\<guillemotleft>prj\<^sub>0 : \<mu>.leg0 \<down>\<down> \<nu>.leg1 \<rightarrow>\<^sub>C \<nu>.apex\<guillemotright>"
using legs_form_cospan by auto
lemma prj_simps [simp]:
shows "C.arr prj\<^sub>1" and "C.dom prj\<^sub>1 = \<mu>.leg0 \<down>\<down> \<nu>.leg1" and "C.cod prj\<^sub>1 = \<mu>.apex"
and "C.arr prj\<^sub>0" and "C.dom prj\<^sub>0 = \<mu>.leg0 \<down>\<down> \<nu>.leg1" and "C.cod prj\<^sub>0 = \<nu>.apex"
using prj_in_hom by auto
sublocale identity_arrow_of_spans C \<open>\<mu> \<star> \<nu>\<close>
using apex_composite dom.ide_apex chine_composite by (unfold_locales, auto)
end
locale three_composable_arrows_of_spans =
span_bicategory +
\<mu>: arrow_of_spans C \<mu> +
\<nu>: arrow_of_spans C \<nu> +
\<pi>: arrow_of_spans C \<pi> +
\<mu>\<nu>: two_composable_arrows_of_spans C prj0 prj1 \<mu> \<nu> +
\<nu>\<pi>: two_composable_arrows_of_spans C prj0 prj1 \<nu> \<pi>
for \<mu> (structure)
and \<nu> (structure)
and \<pi> (structure)
begin
interpretation \<mu>\<nu>\<pi>: arrow_of_spans C \<open>\<mu> \<star> \<nu> \<star> \<pi>\<close>
using \<mu>.arrow_of_spans_axioms \<nu>\<pi>.arrow_of_spans_axioms
arrow_of_spans_hcomp arr_char \<mu>\<nu>.composable \<nu>\<pi>.composable
by force
interpretation \<mu>\<nu>_\<pi>: arrow_of_spans C \<open>(\<mu> \<star> \<nu>) \<star> \<pi>\<close>
using \<mu>\<nu>.arrow_of_spans_axioms \<pi>.arrow_of_spans_axioms
arrow_of_spans_hcomp arr_char \<mu>\<nu>.composable \<nu>\<pi>.composable
by force
lemma composites_are_arrows [simp]:
shows "arr (\<mu> \<star> \<nu> \<star> \<pi>)" and "arr ((\<mu> \<star> \<nu>) \<star> \<pi>)"
using \<mu>\<nu>\<pi>.arrow_of_spans_axioms \<mu>\<nu>_\<pi>.arrow_of_spans_axioms arr_char by auto
lemma composite_in_hom [intro]:
shows "\<guillemotleft>\<mu> \<star> \<nu> \<star> \<pi> : dom \<mu> \<star> dom \<nu> \<star> dom \<pi> \<Rightarrow> cod \<mu> \<star> cod \<nu> \<star> cod \<pi>\<guillemotright>"
and "\<guillemotleft>(\<mu> \<star> \<nu>) \<star> \<pi> : (dom \<mu> \<star> dom \<nu>) \<star> dom \<pi> \<Rightarrow> (cod \<mu> \<star> cod \<nu>) \<star> cod \<pi>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable by auto
lemma composite_simps [simp]:
shows "src (\<mu> \<star> \<nu> \<star> \<pi>) = src \<pi>"
and "src ((\<mu> \<star> \<nu>) \<star> \<pi>) = src \<pi>"
and "trg (\<mu> \<star> \<nu> \<star> \<pi>) = trg \<mu>"
and "trg ((\<mu> \<star> \<nu>) \<star> \<pi>) = trg \<mu>"
and "dom (\<mu> \<star> \<nu> \<star> \<pi>) = dom \<mu> \<star> dom \<nu> \<star> dom \<pi>"
and "dom ((\<mu> \<star> \<nu>) \<star> \<pi>) = (dom \<mu> \<star> dom \<nu>) \<star> dom \<pi>"
and "cod (\<mu> \<star> \<nu> \<star> \<pi>) = cod \<mu> \<star> cod \<nu> \<star> cod \<pi>"
and "cod ((\<mu> \<star> \<nu>) \<star> \<pi>) = (cod \<mu> \<star> cod \<nu>) \<star> cod \<pi>"
by (auto simp add: \<mu>\<nu>.composable \<nu>\<pi>.composable)
lemma chine_composite:
shows "\<mu>\<nu>\<pi>.chine =
\<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]]
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]\<rbrakk>
\<langle>\<nu>.chine \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]
\<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<pi>.dom.leg1]\<rangle> \<cdot>
\<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]]\<rangle>"
and "\<mu>\<nu>_\<pi>.chine =
\<langle>\<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1]
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk>
\<nu>.chine \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1]\<rangle> \<cdot>
\<p>\<^sub>1[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]
\<lbrakk>\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]\<rangle>"
proof -
show "\<mu>\<nu>\<pi>.chine =
\<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]]
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]\<rbrakk>
\<langle>\<nu>.chine \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]
\<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<pi>.dom.leg1]\<rangle> \<cdot>
\<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]]\<rangle>"
unfolding hcomp_def chine_hcomp_def \<mu>\<nu>.composable \<nu>\<pi>.composable
using trg_def \<nu>\<pi>.are_arrows(1-2) \<nu>\<pi>.composable \<nu>\<pi>.composite_is_arrow
\<nu>\<pi>.composite_simps(2) hcomp_def
by (simp add: chine_hcomp_def)
show "\<mu>\<nu>_\<pi>.chine =
\<langle>\<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1]
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk>
\<nu>.chine \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1]\<rangle> \<cdot>
\<p>\<^sub>1[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]
\<lbrakk>\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]\<rangle>"
unfolding hcomp_def chine_hcomp_def \<mu>\<nu>.composable \<nu>\<pi>.composable
using src_def \<mu>\<nu>.are_arrows(1-2) \<mu>\<nu>.composable \<mu>\<nu>.composite_is_arrow
\<mu>\<nu>.composite_simps(1) hcomp_def \<nu>\<pi>.composable
by (simp add: chine_hcomp_def)
qed
end
locale three_composable_identity_arrows_of_spans =
three_composable_arrows_of_spans +
\<mu>: identity_arrow_of_spans C \<mu> +
\<nu>: identity_arrow_of_spans C \<nu> +
\<pi>: identity_arrow_of_spans C \<pi> +
\<mu>\<nu>: two_composable_identity_arrows_of_spans C prj0 prj1 \<mu> \<nu> +
\<nu>\<pi>: two_composable_identity_arrows_of_spans C prj0 prj1 \<nu> \<pi>
begin
lemma composites_are_identities [simp]:
shows "ide (\<mu> \<star> \<nu> \<star> \<pi>)" and "ide ((\<mu> \<star> \<nu>) \<star> \<pi>)"
by (auto simp add: \<mu>\<nu>.composable \<nu>\<pi>.composable)
interpretation \<mu>\<nu>\<pi>: identity_arrow_of_spans C \<open>\<mu> \<star> \<nu> \<star> \<pi>\<close>
using composites_are_identities ide_char' by auto
interpretation \<mu>\<nu>_\<pi>: identity_arrow_of_spans C \<open>(\<mu> \<star> \<nu>) \<star> \<pi>\<close>
using composites_are_identities ide_char' by auto
abbreviation Prj\<^sub>1\<^sub>1
where "Prj\<^sub>1\<^sub>1 \<equiv> \<p>\<^sub>1[\<mu>.leg0, \<nu>.leg1] \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1], \<pi>.leg1]"
abbreviation Prj\<^sub>0\<^sub>1
where "Prj\<^sub>0\<^sub>1 \<equiv> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1] \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1], \<pi>.leg1]"
abbreviation Prj\<^sub>0
where "Prj\<^sub>0 \<equiv> \<p>\<^sub>0[\<nu>.leg0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1], \<pi>.leg1]"
abbreviation Prj\<^sub>1
where "Prj\<^sub>1 \<equiv> \<p>\<^sub>1[\<mu>.leg0, \<nu>.leg1 \<cdot> \<p>\<^sub>1[\<nu>.leg0, \<pi>.leg1]]"
abbreviation Prj\<^sub>1\<^sub>0
where "Prj\<^sub>1\<^sub>0 \<equiv> \<p>\<^sub>1[\<nu>.leg0, \<pi>.leg1] \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<p>\<^sub>1[\<nu>.leg0, \<pi>.leg1]]"
abbreviation Prj\<^sub>0\<^sub>0
where "Prj\<^sub>0\<^sub>0 \<equiv> \<p>\<^sub>0[\<nu>.leg0, \<pi>.leg1] \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<p>\<^sub>1[\<nu>.leg0, \<pi>.leg1]]"
lemma leg0_composite:
shows "\<mu>\<nu>\<pi>.leg0 = \<pi>.leg0 \<cdot> Prj\<^sub>0\<^sub>0"
and "\<mu>\<nu>_\<pi>.leg0 = \<pi>.leg0 \<cdot> Prj\<^sub>0"
proof -
show "\<mu>\<nu>\<pi>.leg0 = \<pi>.leg0 \<cdot> Prj\<^sub>0\<^sub>0"
using hcomp_def \<mu>\<nu>.composable \<nu>\<pi>.composite_is_arrow \<nu>\<pi>.composite_simps(2)
C.comp_assoc
by auto
show "\<mu>\<nu>_\<pi>.leg0 = \<pi>.leg0 \<cdot> Prj\<^sub>0"
using hcomp_def \<mu>\<nu>.composite_is_arrow \<mu>\<nu>.composite_simps(1) \<nu>\<pi>.composable by auto
qed
lemma leg1_composite:
shows "\<mu>\<nu>\<pi>.leg1 = \<mu>.leg1 \<cdot> Prj\<^sub>1"
and "\<mu>\<nu>_\<pi>.leg1 = \<mu>.leg1 \<cdot> Prj\<^sub>1\<^sub>1"
proof -
show "\<mu>\<nu>\<pi>.leg1 = \<mu>.leg1 \<cdot> Prj\<^sub>1"
using hcomp_def \<mu>\<nu>.composable \<nu>\<pi>.composite_is_arrow \<nu>\<pi>.composite_simps(2) by auto
show "\<mu>\<nu>_\<pi>.leg1 = \<mu>.leg1 \<cdot> Prj\<^sub>1\<^sub>1"
using hcomp_def \<mu>\<nu>.composite_is_arrow \<mu>\<nu>.composite_simps(1) \<nu>\<pi>.composable
C.comp_assoc
by auto
qed
definition chine_assoc
where "chine_assoc \<equiv>
\<langle>Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.leg0, \<nu>.leg1 \<cdot> \<p>\<^sub>1[\<nu>.leg0, \<pi>.leg1]\<rbrakk> \<langle>Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>.leg1\<rbrakk> Prj\<^sub>0\<rangle>\<rangle>"
definition chine_assoc'
where "chine_assoc' \<equiv>
\<langle>\<langle>Prj\<^sub>1 \<lbrakk>\<mu>.leg0, \<nu>.leg1\<rbrakk> Prj\<^sub>1\<^sub>0\<rangle> \<lbrakk>\<nu>.leg0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1], \<pi>.leg1\<rbrakk> Prj\<^sub>0\<^sub>0\<rangle>"
(*
* Don't be fooled by how short the following proofs look -- there's a heck of a lot
* going on behind the scenes here!
*)
lemma chine_composite:
shows "\<mu>\<nu>_\<pi>.chine = \<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0 \<down>\<down> \<pi>.leg1"
and "\<mu>\<nu>\<pi>.chine = \<mu>.leg0 \<down>\<down> \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1"
proof -
show "\<mu>\<nu>_\<pi>.chine = \<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0 \<down>\<down> \<pi>.leg1"
using hcomp_def chine_hcomp_arr_ide [of "hcomp \<mu> \<nu>" \<pi>] chine_hcomp_ide_ide
src_def trg_def \<mu>\<nu>.composable \<nu>\<pi>.composable \<mu>\<nu>.ide_composite
\<mu>\<nu>.are_identities \<nu>\<pi>.are_identities(2)
by simp
show "\<mu>\<nu>\<pi>.chine = \<mu>.leg0 \<down>\<down> \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1"
using hcomp_def chine_hcomp_ide_arr [of \<mu> "hcomp \<nu> \<pi>"] chine_hcomp_ide_ide
src_def trg_def \<mu>\<nu>.composable \<nu>\<pi>.composable \<nu>\<pi>.ide_composite
\<mu>\<nu>.are_identities \<nu>\<pi>.are_identities(2)
by simp
qed
lemma prj_in_hom [intro]:
shows "\<guillemotleft>Prj\<^sub>1\<^sub>1 : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<mu>.apex\<guillemotright>"
and "\<guillemotleft>Prj\<^sub>0\<^sub>1 : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<nu>.apex\<guillemotright>"
and "\<guillemotleft>Prj\<^sub>0 : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<pi>.apex\<guillemotright>"
and "\<guillemotleft>Prj\<^sub>1 : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<mu>.apex\<guillemotright>"
and "\<guillemotleft>Prj\<^sub>1\<^sub>0 : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<nu>.apex\<guillemotright>"
and "\<guillemotleft>Prj\<^sub>0\<^sub>0 : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<pi>.apex\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def chine_composite by auto
lemma prj_simps [simp]:
shows "C.arr Prj\<^sub>1\<^sub>1"
and "C.arr Prj\<^sub>0\<^sub>1"
and "C.arr Prj\<^sub>0"
and "C.dom Prj\<^sub>1\<^sub>1 = \<mu>\<nu>_\<pi>.chine"
and "C.dom Prj\<^sub>0\<^sub>1 = \<mu>\<nu>_\<pi>.chine"
and "C.dom Prj\<^sub>0 = \<mu>\<nu>_\<pi>.chine"
and "C.cod Prj\<^sub>1\<^sub>1 = \<mu>.apex"
and "C.cod Prj\<^sub>0\<^sub>1 = \<nu>.apex"
and "C.cod Prj\<^sub>0 = \<pi>.apex"
and "C.arr Prj\<^sub>1"
and "C.arr Prj\<^sub>1\<^sub>0"
and "C.arr Prj\<^sub>0\<^sub>0"
and "C.dom Prj\<^sub>1 = \<mu>\<nu>\<pi>.chine"
and "C.dom Prj\<^sub>1\<^sub>0 = \<mu>\<nu>\<pi>.chine"
and "C.dom Prj\<^sub>0\<^sub>0 = \<mu>\<nu>\<pi>.chine"
and "C.cod Prj\<^sub>1 = \<mu>.apex"
and "C.cod Prj\<^sub>1\<^sub>0 = \<nu>.apex"
and "C.cod Prj\<^sub>0\<^sub>0 = \<pi>.apex"
using prj_in_hom by auto
lemma chine_assoc_props:
shows "\<guillemotleft>chine_assoc : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>\<pi>.chine\<guillemotright>"
and "Prj\<^sub>1 \<cdot> chine_assoc = Prj\<^sub>1\<^sub>1"
and "Prj\<^sub>1\<^sub>0 \<cdot> chine_assoc = Prj\<^sub>0\<^sub>1"
and "Prj\<^sub>0\<^sub>0 \<cdot> chine_assoc = Prj\<^sub>0"
proof -
have 1: "\<nu>.leg0 \<cdot> Prj\<^sub>0\<^sub>1 = \<pi>.leg1 \<cdot> Prj\<^sub>0"
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def
C.pullback_commutes [of "\<nu>.leg0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1]" \<pi>.leg1] C.comp_assoc
by auto
have 2: "\<mu>.leg0 \<cdot> Prj\<^sub>1\<^sub>1 = \<nu>.leg1 \<cdot> Prj\<^sub>0\<^sub>1"
by (metis C.comp_assoc C.pullback_commutes' \<mu>\<nu>.legs_form_cospan(1))
show "\<guillemotleft>chine_assoc : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>\<pi>.chine\<guillemotright>"
unfolding chine_assoc_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def chine_composite C.comp_assoc by auto
show "Prj\<^sub>1 \<cdot> chine_assoc = Prj\<^sub>1\<^sub>1"
unfolding chine_assoc_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def C.comp_assoc
by auto
show "Prj\<^sub>1\<^sub>0 \<cdot> chine_assoc = Prj\<^sub>0\<^sub>1"
unfolding chine_assoc_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def C.comp_assoc
by auto
show "Prj\<^sub>0\<^sub>0 \<cdot> chine_assoc = Prj\<^sub>0"
unfolding chine_assoc_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def C.comp_assoc
by auto
qed
lemma chine_assoc_in_hom [intro]:
shows "\<guillemotleft>chine_assoc : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>\<pi>.chine\<guillemotright>"
using chine_assoc_props(1) by simp
lemma prj_chine_assoc [simp]:
shows "Prj\<^sub>1 \<cdot> chine_assoc = Prj\<^sub>1\<^sub>1"
and "Prj\<^sub>1\<^sub>0 \<cdot> chine_assoc = Prj\<^sub>0\<^sub>1"
and "Prj\<^sub>0\<^sub>0 \<cdot> chine_assoc = Prj\<^sub>0"
using chine_assoc_props(2-4) by auto
lemma chine_assoc'_props:
shows "\<guillemotleft>chine_assoc' : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>_\<pi>.chine\<guillemotright>"
and "Prj\<^sub>1\<^sub>1 \<cdot> chine_assoc' = Prj\<^sub>1"
and "Prj\<^sub>0\<^sub>1 \<cdot> chine_assoc' = Prj\<^sub>1\<^sub>0"
and "Prj\<^sub>0 \<cdot> chine_assoc' = Prj\<^sub>0\<^sub>0"
proof -
have 1: "\<mu>.leg0 \<cdot> Prj\<^sub>1 = \<nu>.leg1 \<cdot> Prj\<^sub>1\<^sub>0"
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable
src_def trg_def C.pullback_commutes [of \<mu>.leg0 "\<nu>.leg1 \<cdot> \<p>\<^sub>1[\<nu>.leg0, \<pi>.leg1]"]
C.comp_assoc
by auto
have 2: "\<nu>.leg0 \<cdot> Prj\<^sub>1\<^sub>0 = \<pi>.leg1 \<cdot> Prj\<^sub>0\<^sub>0"
by (metis C.comp_assoc C.pullback_commutes' \<nu>\<pi>.legs_form_cospan(1))
show "\<guillemotleft>chine_assoc' : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>_\<pi>.chine\<guillemotright>"
unfolding chine_assoc'_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def chine_composite C.comp_assoc
by auto
show "Prj\<^sub>1\<^sub>1 \<cdot> chine_assoc' = Prj\<^sub>1"
unfolding chine_assoc'_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def C.comp_assoc
by auto
show "Prj\<^sub>0\<^sub>1 \<cdot> chine_assoc' = Prj\<^sub>1\<^sub>0"
unfolding chine_assoc'_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def C.comp_assoc
by auto
show "Prj\<^sub>0 \<cdot> chine_assoc' = Prj\<^sub>0\<^sub>0"
unfolding chine_assoc'_def
using \<mu>\<nu>.are_identities \<nu>\<pi>.are_identities \<mu>\<nu>.composable \<nu>\<pi>.composable 1 2
src_def trg_def C.comp_assoc
by auto
qed
lemma chine_assoc'_in_hom [intro]:
shows "\<guillemotleft>chine_assoc' : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>_\<pi>.chine\<guillemotright>"
using chine_assoc'_props(1) by simp
lemma prj_chine_assoc' [simp]:
shows "Prj\<^sub>1\<^sub>1 \<cdot> chine_assoc' = Prj\<^sub>1"
and "Prj\<^sub>0\<^sub>1 \<cdot> chine_assoc' = Prj\<^sub>1\<^sub>0"
and "Prj\<^sub>0 \<cdot> chine_assoc' = Prj\<^sub>0\<^sub>0"
using chine_assoc'_props(2-4) by auto
lemma prj_joint_monic:
assumes "\<guillemotleft>h: a \<rightarrow>\<^sub>C \<mu>\<nu>_\<pi>.chine\<guillemotright>" and "\<guillemotleft>h': a \<rightarrow>\<^sub>C \<mu>\<nu>_\<pi>.chine\<guillemotright>"
and "Prj\<^sub>1\<^sub>1 \<cdot> h = Prj\<^sub>1\<^sub>1 \<cdot> h'" and "Prj\<^sub>0\<^sub>1 \<cdot> h = Prj\<^sub>0\<^sub>1 \<cdot> h'" and "Prj\<^sub>0 \<cdot> h = Prj\<^sub>0 \<cdot> h'"
shows "h = h'"
proof -
have "\<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h = \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h'"
proof -
have "\<mu>\<nu>.prj\<^sub>0 \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h = \<mu>\<nu>.prj\<^sub>0 \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h'"
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def chine_composite(1)
C.comp_assoc
by force
moreover
have "\<mu>\<nu>.prj\<^sub>1 \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h = \<mu>\<nu>.prj\<^sub>1 \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h'"
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def chine_composite(1)
C.comp_assoc
by force
ultimately show ?thesis
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def
chine_composite(1) cod_char
C.prj_joint_monic
[of \<mu>.leg0 \<nu>.leg1 "\<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h"
"\<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>.leg1] \<cdot> h'"]
by auto
qed
moreover have "Prj\<^sub>0 \<cdot> h = Prj\<^sub>0 \<cdot> h'"
using assms cod_char by simp
ultimately show ?thesis
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def
chine_composite(1) cod_char
C.prj_joint_monic [of "\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0" \<pi>.leg1 h h']
by auto
qed
lemma prj'_joint_monic:
assumes "\<guillemotleft>h: a \<rightarrow>\<^sub>C \<mu>\<nu>\<pi>.chine\<guillemotright>" and "\<guillemotleft>h': a \<rightarrow>\<^sub>C \<mu>\<nu>\<pi>.chine\<guillemotright>"
and "Prj\<^sub>1 \<cdot> h = Prj\<^sub>1 \<cdot> h'" and "Prj\<^sub>1\<^sub>0 \<cdot> h = Prj\<^sub>1\<^sub>0 \<cdot> h'" and "Prj\<^sub>0\<^sub>0 \<cdot> h = Prj\<^sub>0\<^sub>0 \<cdot> h'"
shows "h = h'"
proof -
have "\<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h = \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h'"
proof -
have "\<nu>\<pi>.prj\<^sub>0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h = \<nu>\<pi>.prj\<^sub>0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h'"
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def chine_composite(2)
C.comp_assoc
by force
moreover
have "\<nu>\<pi>.prj\<^sub>1 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h = \<nu>\<pi>.prj\<^sub>1 \<cdot> \<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h'"
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def chine_composite(2)
C.comp_assoc
by force
ultimately show ?thesis
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def
chine_composite(2) cod_char
C.prj_joint_monic
[of \<nu>.leg0 \<pi>.leg1 "\<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h"
"\<p>\<^sub>0[\<mu>.leg0, \<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1] \<cdot> h'"]
by auto
qed
moreover have "Prj\<^sub>1 \<cdot> h = Prj\<^sub>1 \<cdot> h'"
using assms cod_char by simp
ultimately show ?thesis
using assms \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def chine_composite(2)
C.prj_joint_monic [of \<mu>.leg0 "\<nu>.leg1 \<cdot> \<nu>\<pi>.prj\<^sub>1" h h']
by auto
qed
lemma chine_assoc_inverse:
shows "C.inverse_arrows chine_assoc chine_assoc'"
proof
show "C.ide (chine_assoc' \<cdot> chine_assoc)"
proof -
have 1: "C.ide \<mu>\<nu>_\<pi>.chine"
using chine_assoc_props(1) C.ide_dom by fastforce
have "chine_assoc' \<cdot> chine_assoc = \<mu>\<nu>_\<pi>.chine"
proof -
have 2: "C.seq chine_assoc' chine_assoc"
using chine_assoc_props(1) chine_assoc'_props(1) by auto
have 3: "C.seq Prj\<^sub>1\<^sub>1 chine_assoc' \<and> C.seq Prj\<^sub>0\<^sub>1 chine_assoc' \<and> C.seq Prj\<^sub>0 chine_assoc'"
using prj_in_hom chine_assoc'_props(1) by auto
have "Prj\<^sub>1\<^sub>1 \<cdot> chine_assoc' \<cdot> chine_assoc = Prj\<^sub>1\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>.chine"
proof -
have "Prj\<^sub>1\<^sub>1 \<cdot> chine_assoc' \<cdot> chine_assoc = (Prj\<^sub>1\<^sub>1 \<cdot> chine_assoc') \<cdot> chine_assoc"
using C.comp_assoc by metis
thus ?thesis using 1 C.comp_arr_dom by simp
qed
moreover have "Prj\<^sub>0\<^sub>1 \<cdot> chine_assoc' \<cdot> chine_assoc = Prj\<^sub>0\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>.chine"
proof -
have "Prj\<^sub>0\<^sub>1 \<cdot> chine_assoc' \<cdot> chine_assoc = (Prj\<^sub>0\<^sub>1 \<cdot> chine_assoc') \<cdot> chine_assoc"
using C.comp_assoc by metis
thus ?thesis using 1 C.comp_arr_dom by simp
qed
moreover have "Prj\<^sub>0 \<cdot> chine_assoc' \<cdot> chine_assoc = Prj\<^sub>0 \<cdot> \<mu>\<nu>_\<pi>.chine"
proof -
have "Prj\<^sub>0 \<cdot> chine_assoc' \<cdot> chine_assoc = (Prj\<^sub>0 \<cdot> chine_assoc') \<cdot> chine_assoc"
using C.comp_assoc by metis
thus ?thesis using 1 C.comp_arr_dom C.comp_arr_ide prj_in_hom(3) by auto
qed
moreover have "\<guillemotleft>\<mu>\<nu>_\<pi>.chine : \<mu>\<nu>_\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>_\<pi>.chine\<guillemotright>"
using chine_assoc_props(1) C.ide_dom [of chine_assoc]
by (elim C.in_homE, auto)
ultimately show ?thesis
using chine_assoc_props(1) chine_assoc'_props(1)
prj_joint_monic [of "chine_assoc' \<cdot> chine_assoc" "\<mu>\<nu>_\<pi>.chine" "\<mu>\<nu>_\<pi>.chine"]
by auto
qed
thus ?thesis
using 1 by simp
qed
show "C.ide (chine_assoc \<cdot> chine_assoc')"
proof -
have 1: "C.ide \<mu>\<nu>\<pi>.chine"
using chine_assoc_props(1) C.ide_cod by fastforce
have "chine_assoc \<cdot> chine_assoc' = \<mu>\<nu>\<pi>.chine"
proof -
have 2: "C.seq chine_assoc chine_assoc'"
using chine_assoc_props(1) chine_assoc'_props(1) by auto
have 3: "C.seq Prj\<^sub>1 chine_assoc \<and> C.seq Prj\<^sub>1\<^sub>0 chine_assoc \<and> C.seq Prj\<^sub>0\<^sub>0 chine_assoc"
using prj_in_hom chine_assoc_props(1) by auto
have "Prj\<^sub>1 \<cdot> chine_assoc \<cdot> chine_assoc' = Prj\<^sub>1 \<cdot> \<mu>\<nu>\<pi>.chine"
proof -
have "Prj\<^sub>1 \<cdot> chine_assoc \<cdot> chine_assoc' = (Prj\<^sub>1 \<cdot> chine_assoc) \<cdot> chine_assoc'"
using C.comp_assoc by metis
thus ?thesis using 1 C.comp_arr_dom prj_in_hom(4) by auto
qed
moreover have "Prj\<^sub>1\<^sub>0 \<cdot> chine_assoc \<cdot> chine_assoc' = Prj\<^sub>1\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine"
proof -
have "Prj\<^sub>1\<^sub>0 \<cdot> chine_assoc \<cdot> chine_assoc' = (Prj\<^sub>1\<^sub>0 \<cdot> chine_assoc) \<cdot> chine_assoc'"
using C.comp_assoc by metis
thus ?thesis using 1 C.comp_arr_dom by simp
qed
moreover have "Prj\<^sub>0\<^sub>0 \<cdot> chine_assoc \<cdot> chine_assoc' = Prj\<^sub>0\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine"
proof -
have "Prj\<^sub>0\<^sub>0 \<cdot> chine_assoc \<cdot> chine_assoc' = (Prj\<^sub>0\<^sub>0 \<cdot> chine_assoc) \<cdot> chine_assoc'"
using C.comp_assoc by metis
thus ?thesis using 1 C.comp_arr_dom by simp
qed
moreover have "\<guillemotleft>\<mu>\<nu>\<pi>.chine : \<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<mu>\<nu>\<pi>.chine\<guillemotright>"
using chine_assoc'_props(1) C.ide_dom [of chine_assoc']
by (elim C.in_homE, auto)
ultimately show ?thesis
using chine_assoc_props(1) chine_assoc'_props(1)
prj'_joint_monic [of "chine_assoc \<cdot> chine_assoc'" "\<mu>\<nu>\<pi>.chine" "\<mu>\<nu>\<pi>.chine"]
by auto
qed
thus ?thesis
using 1 by simp
qed
qed
end
context three_composable_arrows_of_spans
begin
interpretation V: category vcomp
using is_category by auto
interpretation H: horizontal_homs vcomp src trg
using has_horizontal_homs by auto
interpretation dom_\<mu>: arrow_of_spans C \<open>dom \<mu>\<close>
using \<mu>.arrow_of_spans_axioms arr_char [of "dom \<mu>"] by auto
interpretation dom_\<nu>: arrow_of_spans C \<open>dom \<nu>\<close>
using \<nu>.arrow_of_spans_axioms arr_char [of "dom \<nu>"] by auto
interpretation dom_\<pi>: arrow_of_spans C \<open>dom \<pi>\<close>
using \<pi>.arrow_of_spans_axioms arr_char [of "dom \<pi>"] by auto
interpretation doms: three_composable_identity_arrows_of_spans C prj0 prj1
\<open>dom \<mu>\<close> \<open>dom \<nu>\<close> \<open>dom \<pi>\<close>
using \<mu>\<nu>.composable \<nu>\<pi>.composable ide_char [of "dom \<mu>"] ide_char [of "dom \<nu>"]
ide_char [of "dom \<pi>"]
by (unfold_locales, auto)
interpretation cod_\<mu>: arrow_of_spans C \<open>cod \<mu>\<close>
using \<mu>.arrow_of_spans_axioms arr_char [of "cod \<mu>"] by auto
interpretation cod_\<nu>: arrow_of_spans C \<open>cod \<nu>\<close>
using \<nu>.arrow_of_spans_axioms arr_char [of "cod \<nu>"] by auto
interpretation cod_\<pi>: arrow_of_spans C \<open>cod \<pi>\<close>
using \<pi>.arrow_of_spans_axioms arr_char [of "cod \<pi>"] by auto
interpretation cods: three_composable_identity_arrows_of_spans C prj0 prj1
\<open>cod \<mu>\<close> \<open>cod \<nu>\<close> \<open>cod \<pi>\<close>
using \<mu>\<nu>.composable \<nu>\<pi>.composable ide_char [of "cod \<mu>"] ide_char [of "cod \<nu>"]
ide_char [of "cod \<pi>"]
by (unfold_locales, auto)
interpretation \<mu>\<nu>\<pi>: arrow_of_spans C \<open>\<mu> \<star> \<nu> \<star> \<pi>\<close>
using \<mu>.arrow_of_spans_axioms \<nu>\<pi>.arrow_of_spans_axioms
arrow_of_spans_hcomp arr_char \<mu>\<nu>.composable \<nu>\<pi>.composable
by force
interpretation \<mu>\<nu>_\<pi>: arrow_of_spans C \<open>(\<mu> \<star> \<nu>) \<star> \<pi>\<close>
using \<mu>\<nu>.arrow_of_spans_axioms \<pi>.arrow_of_spans_axioms
arrow_of_spans_hcomp arr_char \<mu>\<nu>.composable \<nu>\<pi>.composable
by force
lemma chine_composite':
shows "\<mu>\<nu>\<pi>.chine = \<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]\<rbrakk>
\<langle>\<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 \<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0\<rangle>\<rangle>"
and "\<mu>\<nu>_\<pi>.chine = \<langle>\<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1\<rangle>
\<lbrakk>\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> doms.Prj\<^sub>0\<rangle>"
proof -
show "\<mu>\<nu>_\<pi>.chine = \<langle>\<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1\<rangle>
\<lbrakk>\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> doms.Prj\<^sub>0\<rangle>"
proof -
have "arr (\<mu> \<star> \<nu>)" by simp
thus ?thesis
unfolding hcomp_def chine_hcomp_def dom_char
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char chine_hcomp_props
C.comp_tuple_arr C.pullback_commutes C.comp_assoc
by auto
qed
show "\<mu>\<nu>\<pi>.chine = \<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]\<rbrakk>
\<langle>\<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 \<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0\<rangle>\<rangle>"
proof -
have "arr (\<nu> \<star> \<pi>)" by simp
thus ?thesis
unfolding hcomp_def chine_hcomp_def dom_char
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char chine_hcomp_props
C.comp_tuple_arr C.pullback_commutes C.comp_assoc
by auto
qed
qed
lemma chine_composite_in_hom [intro]:
shows "\<guillemotleft>\<mu>\<nu>_\<pi>.chine : Chn ((dom \<mu> \<star> dom \<nu>) \<star> dom \<pi>) \<rightarrow>\<^sub>C Chn ((cod \<mu> \<star> cod \<nu>) \<star> cod \<pi>)\<guillemotright>"
and "\<guillemotleft>\<mu>\<nu>\<pi>.chine : Chn (dom \<mu> \<star> dom \<nu> \<star> dom \<pi>) \<rightarrow>\<^sub>C Chn (cod \<mu> \<star> cod \<nu> \<star> cod \<pi>)\<guillemotright>"
using Chn_in_hom by auto
lemma cospan_\<mu>\<nu>:
shows "C.cospan \<mu>.dom.leg0 \<nu>.dom.leg1"
using \<mu>\<nu>.legs_form_cospan by simp
lemma cospan_\<nu>\<pi>:
shows "C.cospan \<nu>.dom.leg0 \<pi>.dom.leg1"
using \<nu>\<pi>.legs_form_cospan by simp
lemma commutativities:
shows "\<mu>.cod.leg0 \<cdot> \<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 = \<nu>.cod.leg1 \<cdot> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1"
and "\<pi>.cod.leg1 \<cdot> \<pi>.chine \<cdot> doms.Prj\<^sub>0 =
(\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1]) \<cdot>
\<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1\<rangle>"
proof -
have AB: "\<mu>.dom.leg0 \<cdot> doms.Prj\<^sub>1\<^sub>1 = \<nu>.dom.leg1 \<cdot> doms.Prj\<^sub>0\<^sub>1"
proof -
have "\<mu>.dom.leg0 \<cdot> doms.Prj\<^sub>1\<^sub>1 =
(\<mu>.dom.leg0 \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1]) \<cdot>
\<p>\<^sub>1[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]"
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char C.comp_assoc
by simp
also have "... = (\<nu>.dom.leg1 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1]) \<cdot>
\<p>\<^sub>1[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]"
using C.pullback_commutes' \<mu>\<nu>.legs_form_cospan by auto
also have "... = \<nu>.dom.leg1 \<cdot> doms.Prj\<^sub>0\<^sub>1"
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char C.comp_assoc
by simp
finally show ?thesis by auto
qed
have BC: "\<nu>.dom.leg0 \<cdot> doms.Prj\<^sub>0\<^sub>1 = \<pi>.dom.leg1 \<cdot> doms.Prj\<^sub>0"
proof -
have "\<nu>.dom.leg0 \<cdot> doms.Prj\<^sub>0\<^sub>1 =
(\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1]) \<cdot>
\<p>\<^sub>1[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]"
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char C.comp_assoc
by simp
also have "... = \<pi>.dom.leg1 \<cdot> doms.Prj\<^sub>0"
using C.pullback_commutes' dom_char cod_char \<mu>\<nu>.legs_form_cospan \<nu>\<pi>.legs_form_cospan
by auto
finally show ?thesis by simp
qed
show 1: "\<mu>.cod.leg0 \<cdot> \<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 = \<nu>.cod.leg1 \<cdot> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1"
using AB C.comp_assoc [of \<mu>.cod.leg0 \<mu>.chine]
C.comp_assoc [of \<nu>.cod.leg1 \<nu>.chine doms.Prj\<^sub>0\<^sub>1]
by simp
show "\<pi>.cod.leg1 \<cdot> \<pi>.chine \<cdot> doms.Prj\<^sub>0 =
(\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1]) \<cdot>
\<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1\<rangle>"
proof -
have "(\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1]) \<cdot>
\<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1\<rangle> =
\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1] \<cdot>
\<langle>\<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1 \<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1\<rangle>"
using 1 \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char C.comp_assoc by simp
also have "... = \<nu>.cod.leg0 \<cdot> \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1"
using 1 dom_char \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def by simp
also have "... = (\<nu>.cod.leg0 \<cdot> \<nu>.chine) \<cdot> doms.Prj\<^sub>0\<^sub>1"
using C.comp_assoc [of \<nu>.cod.leg0 \<nu>.chine doms.Prj\<^sub>0\<^sub>1] by simp
also have "... = (\<pi>.cod.leg1 \<cdot> \<pi>.chine) \<cdot> doms.Prj\<^sub>0"
using BC by simp
also have "... = \<pi>.cod.leg1 \<cdot> \<pi>.chine \<cdot> doms.Prj\<^sub>0"
using C.comp_assoc by blast
finally show ?thesis by simp
qed
qed
lemma prj_chine_composite:
shows "cods.Prj\<^sub>1\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>.chine = \<mu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>1"
and "cods.Prj\<^sub>0\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>.chine = \<nu>.chine \<cdot> doms.Prj\<^sub>0\<^sub>1"
and "cods.Prj\<^sub>0 \<cdot> \<mu>\<nu>_\<pi>.chine = \<pi>.chine \<cdot> doms.Prj\<^sub>0"
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char cod_char commutativities
chine_composite' C.comp_assoc
by auto
lemma commutativities':
shows "\<nu>.cod.leg0 \<cdot> \<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 = \<pi>.cod.leg1 \<cdot> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0"
and "\<mu>.cod.leg0 \<cdot> \<mu>.chine \<cdot> doms.Prj\<^sub>1 =
(\<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]) \<cdot>
\<langle>\<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 \<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0\<rangle>"
proof -
have AB: "\<mu>.dom.leg0 \<cdot> doms.Prj\<^sub>1 = \<nu>.dom.leg1 \<cdot> doms.Prj\<^sub>1\<^sub>0"
using C.pullback_commutes' dom_char cod_char \<mu>\<nu>.legs_form_cospan \<nu>\<pi>.legs_form_cospan
C.comp_assoc
by auto
have BC: "\<nu>.dom.leg0 \<cdot> doms.Prj\<^sub>1\<^sub>0 = \<pi>.dom.leg1 \<cdot> doms.Prj\<^sub>0\<^sub>0"
proof -
have "\<nu>.dom.leg0 \<cdot> doms.Prj\<^sub>1\<^sub>0 =
(\<nu>.dom.leg0 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]) \<cdot>
\<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]]"
using dom_char \<mu>\<nu>.legs_form_cospan \<nu>\<pi>.legs_form_cospan C.comp_assoc by simp
also have "... = \<pi>.dom.leg1 \<cdot> doms.Prj\<^sub>0\<^sub>0"
using C.pullback_commutes' dom_char \<mu>\<nu>.legs_form_cospan \<nu>\<pi>.legs_form_cospan C.comp_assoc
by simp
finally show ?thesis by auto
qed
show 1: "\<nu>.cod.leg0 \<cdot> \<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 = \<pi>.cod.leg1 \<cdot> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0"
using BC C.comp_assoc [of \<nu>.cod.leg0 \<nu>.chine doms.Prj\<^sub>1\<^sub>0]
C.comp_assoc [of \<pi>.cod.leg1 \<pi>.chine doms.Prj\<^sub>0\<^sub>0]
doms.prj_in_hom(5-6) dom_char
by force
show "\<mu>.cod.leg0 \<cdot> \<mu>.chine \<cdot> doms.Prj\<^sub>1 =
(\<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]) \<cdot>
\<langle>\<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 \<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0\<rangle>"
proof -
have "(\<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]) \<cdot>
\<langle>\<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 \<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0\<rangle> =
\<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1] \<cdot>
\<langle>\<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0 \<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk> \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0\<rangle>"
using 1 \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char C.comp_assoc by simp
also have "... = \<nu>.cod.leg1 \<cdot> \<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0"
using 1 dom_char \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def
by simp
also have "... = (\<nu>.cod.leg1 \<cdot> \<nu>.chine) \<cdot> doms.Prj\<^sub>1\<^sub>0"
using C.comp_assoc [of \<nu>.cod.leg1 \<nu>.chine doms.Prj\<^sub>1\<^sub>0] by auto
also have "... = (\<mu>.cod.leg0 \<cdot> \<mu>.chine) \<cdot> doms.Prj\<^sub>1"
using AB by simp
also have "... = \<mu>.cod.leg0 \<cdot> \<mu>.chine \<cdot> doms.Prj\<^sub>1"
using C.comp_assoc [of \<mu>.cod.leg0 \<mu>.chine doms.Prj\<^sub>1] doms.prj_in_hom(4) dom_char
by auto
finally show ?thesis by simp
qed
qed
lemma prj'_chine_composite:
shows "cods.Prj\<^sub>1 \<cdot> \<mu>\<nu>\<pi>.chine = \<mu>.chine \<cdot> doms.Prj\<^sub>1"
and "cods.Prj\<^sub>1\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine = \<nu>.chine \<cdot> doms.Prj\<^sub>1\<^sub>0"
and "cods.Prj\<^sub>0\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine = \<pi>.chine \<cdot> doms.Prj\<^sub>0\<^sub>0"
using \<mu>\<nu>.composable \<nu>\<pi>.composable src_def trg_def dom_char cod_char commutativities'
chine_composite' C.comp_assoc
by auto
lemma chine_assoc_naturality:
shows "cods.chine_assoc \<cdot> \<mu>\<nu>_\<pi>.chine = \<mu>\<nu>\<pi>.chine \<cdot> doms.chine_assoc"
proof -
have "\<guillemotleft>cods.chine_assoc \<cdot> \<mu>\<nu>_\<pi>.chine :
Chn ((dom \<mu> \<star> dom \<nu>) \<star> dom \<pi>) \<rightarrow>\<^sub>C Chn (cod \<mu> \<star> cod \<nu> \<star> cod \<pi>)\<guillemotright>"
using cods.chine_assoc_props(1) chine_composite_in_hom(1) by blast
moreover have "\<guillemotleft>\<mu>\<nu>\<pi>.chine \<cdot> doms.chine_assoc :
Chn ((dom \<mu> \<star> dom \<nu>) \<star> dom \<pi>) \<rightarrow>\<^sub>C Chn (cod \<mu> \<star> cod \<nu> \<star> cod \<pi>)\<guillemotright>"
using doms.chine_assoc_props(1) chine_composite_in_hom(2) by blast
moreover
have "cods.Prj\<^sub>1 \<cdot> cods.chine_assoc \<cdot> \<mu>\<nu>_\<pi>.chine =
cods.Prj\<^sub>1 \<cdot> \<mu>\<nu>\<pi>.chine \<cdot> doms.chine_assoc"
using C.comp_assoc doms.chine_assoc_props(2) cods.chine_assoc_props(2)
prj_chine_composite prj'_chine_composite
by metis
moreover have "cods.Prj\<^sub>1\<^sub>0 \<cdot> cods.chine_assoc \<cdot> \<mu>\<nu>_\<pi>.chine =
cods.Prj\<^sub>1\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine \<cdot> doms.chine_assoc"
using C.comp_assoc doms.chine_assoc_props(3) cods.chine_assoc_props(3)
prj_chine_composite prj'_chine_composite
by metis
moreover have "cods.Prj\<^sub>0\<^sub>0 \<cdot> cods.chine_assoc \<cdot> \<mu>\<nu>_\<pi>.chine =
cods.Prj\<^sub>0\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine \<cdot> doms.chine_assoc"
using C.comp_assoc doms.chine_assoc_props(4) cods.chine_assoc_props(4)
prj_chine_composite prj'_chine_composite
by metis
ultimately show ?thesis
using cods.prj'_joint_monic by auto
qed
end
context span_bicategory
begin
abbreviation (input) assoc\<^sub>S\<^sub>B
where "assoc\<^sub>S\<^sub>B f g h \<equiv> \<lparr>Chn = three_composable_identity_arrows_of_spans.chine_assoc
C prj0 prj1 f g h,
Dom = Dom ((f \<star> g) \<star> h), Cod = Cod (f \<star> g \<star> h)\<rparr>"
abbreviation (input) assoc'\<^sub>S\<^sub>B
where "assoc'\<^sub>S\<^sub>B f g h \<equiv> \<lparr>Chn = three_composable_identity_arrows_of_spans.chine_assoc'
C prj0 prj1 f g h,
Dom = Cod (f \<star> g \<star> h), Cod = Dom ((f \<star> g) \<star> h)\<rparr>"
lemma assoc_props:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "src (assoc\<^sub>S\<^sub>B f g h) = src h" and "trg (assoc\<^sub>S\<^sub>B f g h) = trg f"
and "\<guillemotleft>assoc\<^sub>S\<^sub>B f g h : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
and "src (assoc'\<^sub>S\<^sub>B f g h) = src h" and "trg (assoc'\<^sub>S\<^sub>B f g h) = trg f"
and "\<guillemotleft>assoc'\<^sub>S\<^sub>B f g h : f \<star> g \<star> h \<Rightarrow> (f \<star> g) \<star> h\<guillemotright>"
proof -
have fgh: "VVV.ide (f, g, h)"
- using assms VVV.ide_char VV.ide_char VVV.arr_char VV.arr_char by simp
+ using assms VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
interpret f: arrow_of_spans C f
using assms arr_char by fastforce
interpret g: arrow_of_spans C g
using assms arr_char by fastforce
interpret h: arrow_of_spans C h
using assms arr_char by fastforce
interpret fgh: three_composable_arrows_of_spans C prj0 prj1 f g h
using assms arr_char by (unfold_locales, auto)
interpret fgh: three_composable_identity_arrows_of_spans C prj0 prj1 f g h
using assms ide_char by (unfold_locales, auto)
interpret HHfgh: arrow_of_spans C \<open>(f \<star> g) \<star> h\<close>
using assms fgh.composites_are_arrows arrow_of_spans_hcomp by simp
interpret HfHgh: arrow_of_spans C \<open>f \<star> g \<star> h\<close>
using assms fgh.composites_are_arrows arrow_of_spans_hcomp by simp
interpret assoc_fgh: arrow_of_spans C \<open>assoc\<^sub>S\<^sub>B f g h\<close>
apply unfold_locales
apply simp_all
apply (metis C.ideD(2) C.ideD(3) HHfgh.chine_simps(2) HfHgh.chine_simps(3)
fgh.composites_are_identities(1) fgh.composites_are_identities(2)
fgh.chine_assoc_in_hom ide_char)
apply (metis C.comp_assoc fgh.composites_are_identities(1) fgh.leg0_composite(1-2)
fgh.prj_chine_assoc(3) ide_char' identity_arrow_of_spans.cod_simps(2))
by (metis C.comp_assoc fgh.composites_are_identities(1) fgh.leg1_composite(1-2)
fgh.prj_chine_assoc(1) ide_char' identity_arrow_of_spans.cod_simps(3))
interpret assoc'_fgh: arrow_of_spans C \<open>assoc'\<^sub>S\<^sub>B f g h\<close>
apply unfold_locales
apply simp_all
apply (metis C.ideD(2) C.ideD(3) HHfgh.chine_simps(2) HfHgh.chine_simps(3)
fgh.composites_are_identities(1) fgh.composites_are_identities(2)
fgh.chine_assoc'_in_hom ide_char)
using C.comp_assoc fgh.composites_are_identities(1) fgh.leg0_composite(1-2)
ide_char' identity_arrow_of_spans.cod_simps(2)
apply force
using C.comp_assoc fgh.composites_are_identities(1) fgh.leg1_composite(1-2)
fgh.prj_chine_assoc'(1) ide_char' identity_arrow_of_spans.cod_simps(3)
by force
show 1: "\<guillemotleft>assoc\<^sub>S\<^sub>B f g h : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright>"
proof
show 1: "arr (assoc\<^sub>S\<^sub>B f g h)"
using assoc_fgh.arrow_of_spans_axioms arr_char by blast
show "dom (assoc\<^sub>S\<^sub>B f g h) = (f \<star> g) \<star> h"
using fgh 1 dom_char fgh.\<mu>\<nu>.composable fgh.\<nu>\<pi>.composable ideD(2)
by auto
show "cod (assoc\<^sub>S\<^sub>B f g h) = f \<star> g \<star> h"
using fgh 1 HoVH_def cod_char fgh.\<mu>\<nu>.composable fgh.\<nu>\<pi>.composable ideD(3)
by auto
qed
show 2: "\<guillemotleft>assoc'\<^sub>S\<^sub>B f g h : f \<star> g \<star> h \<Rightarrow> (f \<star> g) \<star> h\<guillemotright>"
proof
show 1: "arr (assoc'\<^sub>S\<^sub>B f g h)"
using assoc'_fgh.arrow_of_spans_axioms arr_char by blast
show "dom (assoc'\<^sub>S\<^sub>B f g h) = f \<star> g \<star> h"
using fgh 1 dom_char cod_char ideD(3) by auto
show "cod (assoc'\<^sub>S\<^sub>B f g h) = (f \<star> g) \<star> h"
using fgh 1 cod_char dom_char ideD(2) by auto
qed
show 3: "src (assoc\<^sub>S\<^sub>B f g h) = src h"
proof -
have 4: "src (assoc\<^sub>S\<^sub>B f g h) =
\<lparr>Chn = assoc_fgh.dsrc, Dom = \<lparr>Leg0 = assoc_fgh.dsrc, Leg1 = assoc_fgh.dsrc\<rparr>,
Cod = \<lparr>Leg0 = assoc_fgh.dsrc, Leg1 = assoc_fgh.dsrc\<rparr>\<rparr>"
unfolding src_def using 1 by auto
also have "... = src h"
using fgh.composite_simps(2) src_def by auto
finally show ?thesis by blast
qed
show "src (assoc'\<^sub>S\<^sub>B f g h) = src h"
proof -
have "src (assoc'\<^sub>S\<^sub>B f g h) =
\<lparr>Chn = assoc'_fgh.dsrc, Dom = \<lparr>Leg0 = assoc'_fgh.dsrc, Leg1 = assoc'_fgh.dsrc\<rparr>,
Cod = \<lparr>Leg0 = assoc'_fgh.dsrc, Leg1 = assoc'_fgh.dsrc\<rparr>\<rparr>"
unfolding src_def using 2 by auto
also have "... = src h"
using 1 3 assoc_fgh.cod_src_eq_dom_src arrI src_def by auto
finally show ?thesis by blast
qed
show 4: "trg (assoc\<^sub>S\<^sub>B f g h) = trg f"
proof -
have 5: "trg (assoc\<^sub>S\<^sub>B f g h) =
\<lparr>Chn = assoc_fgh.dtrg, Dom = \<lparr>Leg0 = assoc_fgh.dtrg, Leg1 = assoc_fgh.dtrg\<rparr>,
Cod = \<lparr>Leg0 = assoc_fgh.dtrg, Leg1 = assoc_fgh.dtrg\<rparr>\<rparr>"
unfolding trg_def using 1 by auto
also have "... = trg f"
using fgh.composite_simps(4) trg_def by auto
finally show ?thesis by blast
qed
show "trg (assoc'\<^sub>S\<^sub>B f g h) = trg f"
proof -
have 5: "trg (assoc'\<^sub>S\<^sub>B f g h) =
\<lparr>Chn = assoc'_fgh.dtrg, Dom = \<lparr>Leg0 = assoc'_fgh.dtrg, Leg1 = assoc'_fgh.dtrg\<rparr>,
Cod = \<lparr>Leg0 = assoc'_fgh.dtrg, Leg1 = assoc'_fgh.dtrg\<rparr>\<rparr>"
unfolding trg_def using 2 by auto
also have "... = trg f"
using 1 4 assoc_fgh.cod_trg_eq_dom_trg arrI trg_def by auto
finally show ?thesis by blast
qed
qed
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 "\<guillemotleft>assoc\<^sub>S\<^sub>B f g h : (f \<star> g) \<star> h \<Rightarrow> f \<star> g \<star> h\<guillemotright> "
using assms assoc_props 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 "\<guillemotleft>assoc'\<^sub>S\<^sub>B f g h : f \<star> g \<star> h \<Rightarrow> (f \<star> g) \<star> h\<guillemotright> "
using assms assoc_props by auto
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 (assoc\<^sub>S\<^sub>B f g h)" and "dom (assoc\<^sub>S\<^sub>B f g h) = (f \<star> g) \<star> h"
and "cod (assoc\<^sub>S\<^sub>B f g h) = f \<star> g \<star> h"
and "src (assoc\<^sub>S\<^sub>B f g h) = src h" and "trg (assoc\<^sub>S\<^sub>B f g h) = trg f"
using assms assoc_props(1-3) by (fast, fast, fast, auto)
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 (assoc'\<^sub>S\<^sub>B f g h)" and "dom (assoc'\<^sub>S\<^sub>B f g h) = f \<star> g \<star> h"
and "cod (assoc'\<^sub>S\<^sub>B f g h) = (f \<star> g) \<star> h"
and "src (assoc'\<^sub>S\<^sub>B f g h) = src h" and "trg (assoc'\<^sub>S\<^sub>B f g h) = trg f"
using assms assoc_props(4-6) by (fast, fast, fast, auto)
lemma inverse_assoc_assoc':
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "inverse_arrows (assoc\<^sub>S\<^sub>B f g h) (assoc'\<^sub>S\<^sub>B f g h)"
proof -
interpret f: arrow_of_spans C f using assms arr_char ideD(1) by simp
interpret g: arrow_of_spans C g using assms arr_char ideD(1) by simp
interpret h: arrow_of_spans C h using assms arr_char ideD(1) by simp
interpret fgh: three_composable_arrows_of_spans C prj0 prj1 f g h
using assms arr_char by (unfold_locales, auto)
interpret fgh: three_composable_identity_arrows_of_spans C prj0 prj1 f g h
using assms ide_char
by unfold_locales blast+
interpret afgh: arrow_of_spans C \<open>assoc\<^sub>S\<^sub>B f g h\<close>
using assms assoc_props(3) arr_char by blast
interpret a'fgh: arrow_of_spans C \<open>assoc'\<^sub>S\<^sub>B f g h\<close>
using assms assoc_props(6) arr_char by blast
show ?thesis
proof -
have "inverse_arrows (assoc\<^sub>S\<^sub>B f g h)
\<lparr>Chn = C.inv (Chn (assoc\<^sub>S\<^sub>B f g h)),
Dom = Cod (assoc\<^sub>S\<^sub>B f g h), Cod = Dom (assoc\<^sub>S\<^sub>B f g h)\<rparr>"
using inverse_arrows [of "assoc\<^sub>S\<^sub>B f g h"] afgh.arrow_of_spans_axioms
arr_char fgh.chine_assoc_inverse
by auto
moreover have "C.inv (Chn (assoc\<^sub>S\<^sub>B f g h)) = fgh.chine_assoc'"
using fgh.chine_assoc_inverse C.inv_is_inverse C.inverse_arrow_unique by auto
ultimately show ?thesis by simp
qed
qed
interpretation \<alpha>: transformation_by_components VVV.comp vcomp HoHV HoVH
\<open>\<lambda>fgh. assoc\<^sub>S\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))\<close>
proof
show *: "\<And>fgh. VVV.ide fgh \<Longrightarrow> \<guillemotleft>assoc\<^sub>S\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh)) :
HoHV fgh \<Rightarrow> HoVH fgh\<guillemotright>"
proof -
fix fgh
assume fgh: "VVV.ide fgh"
show "\<guillemotleft>assoc\<^sub>S\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh)) : HoHV fgh \<Rightarrow> HoVH fgh\<guillemotright>"
unfolding HoHV_def HoVH_def
using fgh assoc_in_hom [of "fst fgh" "fst (snd fgh)" "snd (snd fgh)"]
- VVV.arr_char VVV.ide_char VV.arr_char
+ VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
qed
show "\<And>\<mu>\<nu>\<pi>. VVV.arr \<mu>\<nu>\<pi> \<Longrightarrow>
assoc\<^sub>S\<^sub>B (fst (VVV.cod \<mu>\<nu>\<pi>)) (fst (snd (VVV.cod \<mu>\<nu>\<pi>)))
(snd (snd (VVV.cod \<mu>\<nu>\<pi>))) \<bullet>
HoHV \<mu>\<nu>\<pi> =
HoVH \<mu>\<nu>\<pi> \<bullet> assoc\<^sub>S\<^sub>B (fst (VVV.dom \<mu>\<nu>\<pi>)) (fst (snd (VVV.dom \<mu>\<nu>\<pi>)))
(snd (snd (VVV.dom \<mu>\<nu>\<pi>)))"
proof -
fix \<mu>\<nu>\<pi>
assume \<mu>\<nu>\<pi>: "VVV.arr \<mu>\<nu>\<pi>"
interpret \<mu>: arrow_of_spans C \<open>fst \<mu>\<nu>\<pi>\<close>
- using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char arr_char by auto
+ using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by auto
interpret \<nu>: arrow_of_spans C \<open>fst (snd \<mu>\<nu>\<pi>)\<close>
- using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char VV.arr_char arr_char by auto
+ using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by auto
interpret \<pi>: arrow_of_spans C \<open>snd (snd \<mu>\<nu>\<pi>)\<close>
- using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char VV.arr_char arr_char by auto
+ using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by auto
interpret \<mu>\<nu>\<pi>: three_composable_arrows_of_spans C prj0 prj1
\<open>fst \<mu>\<nu>\<pi>\<close> \<open>fst (snd \<mu>\<nu>\<pi>)\<close> \<open>snd (snd \<mu>\<nu>\<pi>)\<close>
- using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char VV.arr_char arr_char
+ using \<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char
by (unfold_locales, auto)
interpret HoHV_\<mu>\<nu>\<pi>: arrow_of_spans C \<open>(fst \<mu>\<nu>\<pi> \<star> fst (snd \<mu>\<nu>\<pi>)) \<star> snd (snd \<mu>\<nu>\<pi>)\<close>
proof -
have "arr (HoHV \<mu>\<nu>\<pi>)"
using \<mu>\<nu>\<pi> by simp
thus "arrow_of_spans C ((fst \<mu>\<nu>\<pi> \<star> fst (snd \<mu>\<nu>\<pi>)) \<star> snd (snd \<mu>\<nu>\<pi>))"
using \<mu>\<nu>\<pi> HoHV_def arr_char by auto
qed
interpret HoVH_\<mu>\<nu>\<pi>: arrow_of_spans C \<open>fst \<mu>\<nu>\<pi> \<star> fst (snd \<mu>\<nu>\<pi>) \<star> snd (snd \<mu>\<nu>\<pi>)\<close>
proof -
have "arr (HoVH \<mu>\<nu>\<pi>)"
using \<mu>\<nu>\<pi> by simp
thus "arrow_of_spans C (fst \<mu>\<nu>\<pi> \<star> fst (snd \<mu>\<nu>\<pi>) \<star> snd (snd \<mu>\<nu>\<pi>))"
using \<mu>\<nu>\<pi> HoVH_def arr_char by auto
qed
have dom_\<mu>\<nu>\<pi>: "VVV.ide (VVV.dom \<mu>\<nu>\<pi>)"
using \<mu>\<nu>\<pi> VVV.ide_dom by blast
interpret dom_\<mu>: identity_arrow_of_spans C \<open>fst (VVV.dom \<mu>\<nu>\<pi>)\<close>
- using dom_\<mu>\<nu>\<pi> VVV.ide_char VV.ide_char ide_char' by blast
+ using dom_\<mu>\<nu>\<pi> VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' by blast
interpret dom_\<nu>: identity_arrow_of_spans C \<open>fst (snd (VVV.dom \<mu>\<nu>\<pi>))\<close>
- using dom_\<mu>\<nu>\<pi> VVV.ide_char VV.ide_char ide_char' by blast
+ using dom_\<mu>\<nu>\<pi> VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' by blast
interpret dom_\<pi>: identity_arrow_of_spans C \<open>snd (snd (VVV.dom \<mu>\<nu>\<pi>))\<close>
- using dom_\<mu>\<nu>\<pi> VVV.ide_char VV.ide_char ide_char' by blast
+ using dom_\<mu>\<nu>\<pi> VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' by blast
interpret dom_\<mu>\<nu>\<pi>: three_composable_identity_arrows_of_spans C prj0 prj1
\<open>fst (VVV.dom \<mu>\<nu>\<pi>)\<close> \<open>fst (snd (VVV.dom \<mu>\<nu>\<pi>))\<close>
\<open>snd (snd (VVV.dom \<mu>\<nu>\<pi>))\<close>
- using dom_\<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char VV.arr_char
+ using dom_\<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (unfold_locales, auto)
interpret assoc_dom_\<mu>\<nu>\<pi>: arrow_of_spans C
\<open>assoc\<^sub>S\<^sub>B (fst (VVV.dom \<mu>\<nu>\<pi>)) (fst (snd (VVV.dom \<mu>\<nu>\<pi>)))
(snd (snd (VVV.dom \<mu>\<nu>\<pi>)))\<close>
using \<mu>\<nu>\<pi> VVV.ide_dom * arr_char by fast
have cod_\<mu>\<nu>\<pi>: "VVV.ide (VVV.cod \<mu>\<nu>\<pi>)"
using \<mu>\<nu>\<pi> VVV.ide_cod by blast
interpret cod_\<mu>: identity_arrow_of_spans C \<open>fst (VVV.cod \<mu>\<nu>\<pi>)\<close>
- using cod_\<mu>\<nu>\<pi> VVV.ide_char VV.ide_char ide_char' by blast
+ using cod_\<mu>\<nu>\<pi> VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' by blast
interpret cod_\<nu>: identity_arrow_of_spans C \<open>fst (snd (VVV.cod \<mu>\<nu>\<pi>))\<close>
- using cod_\<mu>\<nu>\<pi> VVV.ide_char VV.ide_char ide_char' by blast
+ using cod_\<mu>\<nu>\<pi> VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' by blast
interpret cod_\<pi>: identity_arrow_of_spans C \<open>snd (snd (VVV.cod \<mu>\<nu>\<pi>))\<close>
- using cod_\<mu>\<nu>\<pi> VVV.ide_char VV.ide_char ide_char' by blast
+ using cod_\<mu>\<nu>\<pi> VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' by blast
interpret cod_\<mu>\<nu>\<pi>: three_composable_identity_arrows_of_spans C prj0 prj1
\<open>fst (VVV.cod \<mu>\<nu>\<pi>)\<close> \<open>fst (snd (VVV.cod \<mu>\<nu>\<pi>))\<close>
\<open>snd (snd (VVV.cod \<mu>\<nu>\<pi>))\<close>
- using cod_\<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char VV.arr_char
+ using cod_\<mu>\<nu>\<pi> VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (unfold_locales, auto)
interpret assoc_cod_\<mu>\<nu>\<pi>: arrow_of_spans C
\<open>assoc\<^sub>S\<^sub>B (fst (VVV.cod \<mu>\<nu>\<pi>)) (fst (snd (VVV.cod \<mu>\<nu>\<pi>)))
(snd (snd (VVV.cod \<mu>\<nu>\<pi>)))\<close>
using \<mu>\<nu>\<pi> VVV.ide_cod * arr_char by fast
have dom_legs:
"dom_\<mu>.leg0 = \<mu>.dom.leg0 \<and> dom_\<nu>.leg0 = \<nu>.dom.leg0 \<and> dom_\<pi>.leg0 = \<pi>.dom.leg0 \<and>
dom_\<mu>.leg1 = \<mu>.dom.leg1 \<and> dom_\<nu>.leg1 = \<nu>.dom.leg1 \<and> dom_\<pi>.leg1 = \<pi>.dom.leg1"
- using VVV.arr_char VVV.dom_char dom_char \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.dom_char\<^sub>S\<^sub>b\<^sub>C dom_char \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable
by auto
have cod_legs:
"cod_\<mu>.leg0 = \<mu>.cod.leg0 \<and> cod_\<nu>.leg0 = \<nu>.cod.leg0 \<and> cod_\<pi>.leg0 = \<pi>.cod.leg0 \<and>
cod_\<mu>.leg1 = \<mu>.cod.leg1 \<and> cod_\<nu>.leg1 = \<nu>.cod.leg1 \<and> cod_\<pi>.leg1 = \<pi>.cod.leg1"
- using \<mu>\<nu>\<pi> VVV.cod_char cod_char by auto
+ using \<mu>\<nu>\<pi> VVV.cod_char\<^sub>S\<^sub>b\<^sub>C cod_char by auto
have Prj\<^sub>1\<^sub>1_dom: "dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1 =
\<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1] \<cdot>
\<p>\<^sub>1[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]"
using dom_legs by argo
have Prj\<^sub>1\<^sub>1_cod: "cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1 =
\<p>\<^sub>1[\<mu>.cod.leg0, \<nu>.cod.leg1] \<cdot>
\<p>\<^sub>1[\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1]"
using cod_legs by argo
have Prj\<^sub>0_dom: "dom_\<mu>\<nu>\<pi>.Prj\<^sub>0 = \<p>\<^sub>0[\<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1], \<pi>.dom.leg1]"
using dom_legs by argo
have Prj\<^sub>0_cod: "cod_\<mu>\<nu>\<pi>.Prj\<^sub>0 = \<p>\<^sub>0[\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1]"
using cod_legs by argo
have Dom: "Dom ((fst (VVV.dom \<mu>\<nu>\<pi>) \<star> fst (snd (VVV.dom \<mu>\<nu>\<pi>))) \<star>
snd (snd (VVV.dom \<mu>\<nu>\<pi>))) =
\<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>"
using dom_\<mu>\<nu>\<pi>.leg0_composite(2) dom_\<mu>\<nu>\<pi>.leg1_composite(2) dom_legs by auto
have Cod: "Cod (fst (VVV.dom \<mu>\<nu>\<pi>) \<star> fst (snd (VVV.dom \<mu>\<nu>\<pi>)) \<star>
snd (snd (VVV.dom \<mu>\<nu>\<pi>))) =
\<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>"
proof -
have "arr (dom (fst (snd \<mu>\<nu>\<pi>)) \<star> dom (snd (snd \<mu>\<nu>\<pi>)))"
using \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable by simp
thus ?thesis
using \<mu>\<nu>\<pi> hcomp_def dom_legs ide_dom dom_char
apply simp
using \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable src_def trg_def dom_char C.comp_assoc
VVV.dom_simp VV.dom_simp
by auto
qed
have Dom': "Dom ((fst (VVV.cod \<mu>\<nu>\<pi>) \<star> fst (snd (VVV.cod \<mu>\<nu>\<pi>))) \<star>
snd (snd (VVV.cod \<mu>\<nu>\<pi>))) =
\<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>"
by (simp add: cod_\<mu>\<nu>\<pi>.leg0_composite(2) cod_\<mu>\<nu>\<pi>.leg1_composite(2) cod_legs)
have Cod': "Cod (fst (VVV.cod \<mu>\<nu>\<pi>) \<star> fst (snd (VVV.cod \<mu>\<nu>\<pi>)) \<star>
snd (snd (VVV.cod \<mu>\<nu>\<pi>))) =
\<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>"
proof -
have "arr (cod (fst (snd \<mu>\<nu>\<pi>)) \<star> cod (snd (snd \<mu>\<nu>\<pi>)))"
using \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable by simp
moreover have "\<mu>.dsrc = \<nu>.dtrg"
using \<mu>\<nu>\<pi>.\<mu>\<nu>.composable src_def trg_def cod_char by simp
ultimately show ?thesis
using \<mu>\<nu>\<pi> hcomp_def cod_legs ide_cod cod_char VVV.cod_simp VV.cod_simp
apply simp
using \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable src_def trg_def cod_char C.comp_assoc
by auto
qed
have assoc_dom:
"assoc\<^sub>S\<^sub>B (fst (VVV.dom \<mu>\<nu>\<pi>)) (fst (snd (VVV.dom \<mu>\<nu>\<pi>)))
(snd (snd (VVV.dom \<mu>\<nu>\<pi>))) =
\<lparr>Chn = dom_\<mu>\<nu>\<pi>.chine_assoc,
Dom = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>\<rparr>"
using Dom Cod by simp
have assoc_cod:
"assoc\<^sub>S\<^sub>B (fst (VVV.cod \<mu>\<nu>\<pi>)) (fst (snd (VVV.cod \<mu>\<nu>\<pi>)))
(snd (snd (VVV.cod \<mu>\<nu>\<pi>))) =
\<lparr>Chn = cod_\<mu>\<nu>\<pi>.chine_assoc,
Dom = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>\<rparr>"
using Dom' Cod' by simp
have HoHV_\<mu>\<nu>\<pi>:
"HoHV \<mu>\<nu>\<pi> =
\<lparr>Chn = HoHV_\<mu>\<nu>\<pi>.chine,
Dom = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>\<rparr>"
proof -
have "arr \<lparr>Chn = \<langle>\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1]
\<lbrakk>\<mu>.cod.leg0, \<nu>.cod.leg1\<rbrakk>
\<nu>.chine \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1]\<rangle>,
Dom = \<lparr>Leg0 = \<nu>.dom.leg0 \<cdot> \<p>\<^sub>0[\<mu>.dom.leg0, \<nu>.dom.leg1],
Leg1 = \<mu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<mu>.dom.leg0, \<nu>.dom.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1],
Leg1 = \<mu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<mu>.cod.leg0, \<nu>.cod.leg1]\<rparr>\<rparr>"
unfolding hcomp_def chine_hcomp_def
using \<mu>\<nu>\<pi> hcomp_def dom_legs cod_legs ide_dom ide_cod dom_char cod_char
\<mu>\<nu>\<pi>.\<mu>\<nu>.composable
by (metis (no_types, lifting) hseq_char(1) \<mu>\<nu>\<pi>.\<mu>\<nu>.composite_is_arrow chine_hcomp_def)
moreover have "(\<mu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<mu>.cod.leg0, \<nu>.cod.leg1]) \<cdot>
\<p>\<^sub>1[\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1] =
\<mu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<mu>.cod.leg0, \<nu>.cod.leg1] \<cdot>
\<p>\<^sub>1[\<nu>.cod.leg0 \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1], \<pi>.cod.leg1]"
using C.comp_assoc by simp
ultimately show ?thesis
unfolding HoHV_def hcomp_def chine_hcomp_def
using \<mu>\<nu>\<pi> \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable src_def trg_def dom_legs cod_legs
C.comp_assoc
by simp
qed
have HoVH_\<mu>\<nu>\<pi>:
"HoVH \<mu>\<nu>\<pi> =
\<lparr>Chn = HoVH_\<mu>\<nu>\<pi>.chine,
Dom = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>\<rparr>"
proof -
have "arr \<lparr>Chn = \<langle>\<nu>.chine \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]
\<lbrakk>\<nu>.cod.leg0, \<pi>.cod.leg1\<rbrakk>
\<pi>.chine \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<pi>.dom.leg1]\<rangle>,
Dom = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> \<p>\<^sub>0[\<nu>.dom.leg0, \<pi>.dom.leg1],
Leg1 = \<nu>.dom.leg1 \<cdot> \<p>\<^sub>1[\<nu>.dom.leg0, \<pi>.dom.leg1]\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<pi>.cod.leg1],
Leg1 = \<nu>.cod.leg1 \<cdot> \<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]\<rparr>\<rparr>"
unfolding hcomp_def chine_hcomp_def
using \<mu>\<nu>\<pi> hcomp_def dom_legs cod_legs ide_dom ide_cod dom_char cod_char
\<mu>\<nu>\<pi>.\<nu>\<pi>.composable
by (metis (no_types, lifting) hseq_char \<mu>\<nu>\<pi>.\<nu>\<pi>.composite_is_arrow chine_hcomp_def)
moreover have "(\<pi>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<pi>.cod.leg1]) \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1 \<cdot>
\<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]] =
\<pi>.cod.leg0 \<cdot> \<p>\<^sub>0[\<nu>.cod.leg0, \<pi>.cod.leg1] \<cdot> \<p>\<^sub>0[\<mu>.cod.leg0, \<nu>.cod.leg1 \<cdot>
\<p>\<^sub>1[\<nu>.cod.leg0, \<pi>.cod.leg1]]"
using C.comp_assoc by simp
ultimately show ?thesis
unfolding HoVH_def hcomp_def chine_hcomp_def
using \<mu>\<nu>\<pi> \<mu>\<nu>\<pi>.\<mu>\<nu>.composable \<mu>\<nu>\<pi>.\<nu>\<pi>.composable src_def trg_def dom_legs cod_legs
C.comp_assoc
by simp
qed
have "assoc\<^sub>S\<^sub>B (fst (VVV.cod \<mu>\<nu>\<pi>)) (fst (snd (VVV.cod \<mu>\<nu>\<pi>)))
(snd (snd (VVV.cod \<mu>\<nu>\<pi>))) \<bullet>
HoHV \<mu>\<nu>\<pi> =
\<lparr>Chn = cod_\<mu>\<nu>\<pi>.chine_assoc \<cdot> HoHV_\<mu>\<nu>\<pi>.chine,
Dom = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>\<rparr>"
proof -
have "arr (HoHV \<mu>\<nu>\<pi>)"
using \<mu>\<nu>\<pi> by simp
thus ?thesis
using vcomp_def HoHV_\<mu>\<nu>\<pi> HoHV_\<mu>\<nu>\<pi>.arrow_of_spans_axioms
assoc_cod_\<mu>\<nu>\<pi>.arrow_of_spans_axioms assoc_cod dom_legs cod_legs
arr_char
by simp
qed
moreover
have "HoVH \<mu>\<nu>\<pi> \<bullet>
assoc\<^sub>S\<^sub>B (fst (VVV.dom \<mu>\<nu>\<pi>)) (fst (snd (VVV.dom \<mu>\<nu>\<pi>)))
(snd (snd (VVV.dom \<mu>\<nu>\<pi>))) =
\<lparr>Chn = HoVH_\<mu>\<nu>\<pi>.chine \<cdot> dom_\<mu>\<nu>\<pi>.chine_assoc,
Dom = \<lparr>Leg0 = \<pi>.dom.leg0 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>0, Leg1 = \<mu>.dom.leg1 \<cdot> dom_\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = \<pi>.cod.leg0 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0, Leg1 = \<mu>.cod.leg1 \<cdot> cod_\<mu>\<nu>\<pi>.Prj\<^sub>1\<rparr>\<rparr>"
proof -
have "arr (HoVH \<mu>\<nu>\<pi>)"
using \<mu>\<nu>\<pi> by simp
thus ?thesis
using vcomp_def HoVH_\<mu>\<nu>\<pi> HoVH_\<mu>\<nu>\<pi>.arrow_of_spans_axioms
assoc_dom_\<mu>\<nu>\<pi>.arrow_of_spans_axioms assoc_dom dom_legs cod_legs
arr_char
by simp
qed
moreover
have "cod_\<mu>\<nu>\<pi>.chine_assoc \<cdot> HoHV_\<mu>\<nu>\<pi>.chine = HoVH_\<mu>\<nu>\<pi>.chine \<cdot> dom_\<mu>\<nu>\<pi>.chine_assoc"
using \<mu>\<nu>\<pi> HoHV_def HoVH_def \<mu>\<nu>\<pi>.chine_assoc_naturality
VVV.dom_simp VV.dom_simp VVV.cod_simp VV.cod_simp by simp
ultimately show "assoc\<^sub>S\<^sub>B (fst (VVV.cod \<mu>\<nu>\<pi>)) (fst (snd (VVV.cod \<mu>\<nu>\<pi>)))
(snd (snd (VVV.cod \<mu>\<nu>\<pi>))) \<bullet>
HoHV \<mu>\<nu>\<pi> =
HoVH \<mu>\<nu>\<pi> \<bullet>
assoc\<^sub>S\<^sub>B (fst (VVV.dom \<mu>\<nu>\<pi>)) (fst (snd (VVV.dom \<mu>\<nu>\<pi>)))
(snd (snd (VVV.dom \<mu>\<nu>\<pi>)))"
by argo
qed
qed
definition assoc ("\<a>[_, _, _]")
where "assoc \<equiv> \<lambda>\<mu> \<nu> \<pi>. \<alpha>.map (\<mu>, \<nu>, \<pi>)"
abbreviation (input) \<alpha>\<^sub>S\<^sub>B
where "\<alpha>\<^sub>S\<^sub>B \<equiv> \<lambda>\<mu>\<nu>\<pi>. assoc (fst \<mu>\<nu>\<pi>) (fst (snd \<mu>\<nu>\<pi>)) (snd (snd \<mu>\<nu>\<pi>))"
lemma \<alpha>_ide:
assumes "ide f" and "ide g" and "ide h"
and "src f = trg g" and "src g = trg h"
shows "\<alpha>\<^sub>S\<^sub>B (f, g, h) = assoc\<^sub>S\<^sub>B f g h"
- using assms assoc_def \<alpha>.map_simp_ide VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char
+ using assms assoc_def \<alpha>.map_simp_ide VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
lemma natural_transformation_\<alpha>:
shows "natural_transformation VVV.comp vcomp HoHV HoVH \<alpha>\<^sub>S\<^sub>B"
using assoc_def \<alpha>.natural_transformation_axioms by auto
interpretation \<alpha>: natural_transformation VVV.comp vcomp HoHV HoVH \<alpha>\<^sub>S\<^sub>B
using natural_transformation_\<alpha> by simp
sublocale \<alpha>: natural_isomorphism VVV.comp vcomp HoHV HoVH \<alpha>\<^sub>S\<^sub>B
proof
show "\<And>fgh. VVV.ide fgh \<Longrightarrow> iso (\<alpha>\<^sub>S\<^sub>B fgh)"
proof -
fix fgh
assume fgh: "VVV.ide fgh"
interpret f: arrow_of_spans C \<open>fst fgh\<close>
- using fgh VVV.ide_char VVV.arr_char arr_char by auto
+ using fgh VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by auto
interpret g: arrow_of_spans C \<open>fst (snd fgh)\<close>
- using fgh VVV.ide_char VVV.arr_char VV.arr_char arr_char by auto
+ using fgh VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by auto
interpret h: arrow_of_spans C \<open>snd (snd fgh)\<close>
- using fgh VVV.ide_char VVV.arr_char VV.arr_char arr_char by auto
+ using fgh VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by auto
interpret fgh: three_composable_arrows_of_spans C prj0 prj1
\<open>fst fgh\<close> \<open>fst (snd fgh)\<close> \<open>snd (snd fgh)\<close>
- using fgh VVV.ide_char VVV.arr_char VV.arr_char arr_char
+ using fgh VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char
by (unfold_locales, auto)
interpret fgh: three_composable_identity_arrows_of_spans C prj0 prj1
\<open>fst fgh\<close> \<open>fst (snd fgh)\<close> \<open>snd (snd fgh)\<close>
- using fgh VVV.ide_char VV.ide_char ide_char
+ using fgh VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char
by unfold_locales blast+
have 1: "arr (\<alpha>\<^sub>S\<^sub>B fgh)"
using fgh \<alpha>.preserves_reflects_arr VVV.ideD(1) by blast
have 2: "\<alpha>\<^sub>S\<^sub>B fgh = assoc\<^sub>S\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))"
using fgh assoc_def \<alpha>_ide [of "fst fgh" "fst (snd fgh)" "snd (snd fgh)"]
- VVV.ide_char VV.ide_char VVV.arr_char VV.arr_char
+ VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "iso ..."
using 1 2 iso_char [of "assoc\<^sub>S\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))"]
fgh.chine_assoc_inverse by auto
ultimately show "iso (\<alpha>\<^sub>S\<^sub>B fgh)" by argo
qed
qed
lemma natural_isomorphism_\<alpha>:
shows "natural_isomorphism VVV.comp vcomp HoHV HoVH \<alpha>\<^sub>S\<^sub>B"
..
end
locale four_composable_arrows_of_spans =
span_bicategory +
\<mu>: arrow_of_spans C \<mu> +
\<nu>: arrow_of_spans C \<nu> +
\<pi>: arrow_of_spans C \<pi> +
\<rho>: arrow_of_spans C \<rho> +
\<mu>\<nu>: two_composable_arrows_of_spans C prj0 prj1 \<mu> \<nu> +
\<nu>\<pi>: two_composable_arrows_of_spans C prj0 prj1 \<nu> \<pi> +
\<pi>\<rho>: two_composable_arrows_of_spans C prj0 prj1 \<pi> \<rho> +
\<mu>\<nu>\<pi>: three_composable_arrows_of_spans C prj0 prj1 \<mu> \<nu> \<pi> +
\<nu>\<pi>\<rho>: three_composable_arrows_of_spans C prj0 prj1 \<nu> \<pi> \<rho>
for \<mu> (structure)
and \<nu> (structure)
and \<pi> (structure)
and \<rho> (structure)
locale four_composable_identity_arrows_of_spans =
four_composable_arrows_of_spans +
\<mu>: identity_arrow_of_spans C \<mu> +
\<nu>: identity_arrow_of_spans C \<nu> +
\<pi>: identity_arrow_of_spans C \<pi> +
\<rho>: identity_arrow_of_spans C \<rho> +
\<mu>\<nu>: two_composable_identity_arrows_of_spans C prj0 prj1 \<mu> \<nu> +
\<nu>\<pi>: two_composable_identity_arrows_of_spans C prj0 prj1 \<nu> \<pi> +
\<pi>\<rho>: two_composable_identity_arrows_of_spans C prj0 prj1 \<pi> \<rho> +
\<mu>\<nu>\<pi>: three_composable_identity_arrows_of_spans C prj0 prj1 \<mu> \<nu> \<pi> +
\<nu>\<pi>\<rho>: three_composable_identity_arrows_of_spans C prj0 prj1 \<nu> \<pi> \<rho>
begin
interpretation H: horizontal_composition vcomp hcomp src trg
using has_horizontal_composition by auto
text \<open>
The following interpretations provide us with some systematic names
for a lot of things.
\<close>
interpretation H\<mu>H\<nu>\<pi>: identity_arrow_of_spans C \<open>\<mu> \<star> \<nu> \<star> \<pi>\<close>
using \<mu>\<nu>\<pi>.composites_are_identities ide_char' by auto
interpretation HH\<mu>\<nu>\<pi>: identity_arrow_of_spans C \<open>(\<mu> \<star> \<nu>) \<star> \<pi>\<close>
using \<mu>\<nu>\<pi>.composites_are_identities ide_char' by auto
interpretation H\<nu>H\<pi>\<rho>: identity_arrow_of_spans C \<open>\<nu> \<star> \<pi> \<star> \<rho>\<close>
using \<nu>\<pi>\<rho>.composites_are_identities ide_char' by auto
interpretation HH\<nu>\<pi>\<rho>: identity_arrow_of_spans C \<open>(\<nu> \<star> \<pi>) \<star> \<rho>\<close>
using \<nu>\<pi>\<rho>.composites_are_identities ide_char' by auto
interpretation H\<mu>H\<nu>H\<pi>\<rho>: arrow_of_spans C \<open>\<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<close>
using arrow_of_spans_hcomp \<mu>\<nu>.composable \<nu>\<pi>\<rho>.composites_are_arrows(1) by auto
interpretation H\<mu>HH\<nu>\<pi>\<rho>: arrow_of_spans C \<open>\<mu> \<star> (\<nu> \<star> \<pi>) \<star> \<rho>\<close>
using arrow_of_spans_hcomp \<mu>\<nu>.composable \<nu>\<pi>\<rho>.composites_are_arrows(1) by auto
interpretation HH\<mu>\<nu>H\<pi>\<rho>: arrow_of_spans C \<open>(\<mu> \<star> \<nu>) \<star> \<pi> \<star> \<rho>\<close>
using hseq_char' match_4 \<mu>\<nu>\<pi>.composites_are_arrows(2) \<pi>\<rho>.composite_is_arrow arr_char
by auto
interpretation HH\<mu>H\<nu>\<pi>\<rho>: arrow_of_spans C \<open>(\<mu> \<star> \<nu> \<star> \<pi>) \<star> \<rho>\<close>
using arrow_of_spans_hcomp \<pi>\<rho>.composable \<mu>\<nu>\<pi>.composites_are_arrows(1) by auto
interpretation HHH\<mu>\<nu>\<pi>\<rho>: arrow_of_spans C \<open>((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho>\<close>
using arrow_of_spans_hcomp \<pi>\<rho>.composable \<mu>\<nu>\<pi>.composites_are_arrows(2) by auto
interpretation assoc\<mu>\<nu>\<pi>: arrow_of_spans C \<open>assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>\<close>
using arr_char \<mu>\<nu>.composable \<nu>\<pi>.composable assoc_simps(1) by auto
interpretation assoc\<nu>\<pi>\<rho>: arrow_of_spans C \<open>assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>\<close>
using arr_char \<nu>\<pi>.composable \<pi>\<rho>.composable assoc_simps(1) by auto
interpretation \<mu>_\<nu>\<pi>: two_composable_identity_arrows_of_spans C prj0 prj1 \<mu> \<open>\<nu> \<star> \<pi>\<close>
using \<mu>\<nu>.composable \<nu>\<pi>.composable by (unfold_locales, auto)
interpretation \<mu>\<nu>_\<pi>: two_composable_identity_arrows_of_spans C prj0 prj1 \<open>\<mu> \<star> \<nu>\<close> \<pi>
using \<mu>\<nu>.composable \<nu>\<pi>.composable by (unfold_locales, auto)
interpretation \<nu>_\<pi>\<rho>: two_composable_identity_arrows_of_spans C prj0 prj1 \<nu> \<open>\<pi> \<star> \<rho>\<close>
using \<nu>\<pi>.composable \<pi>\<rho>.composable by (unfold_locales, auto)
interpretation \<nu>\<pi>_\<rho>: two_composable_identity_arrows_of_spans C prj0 prj1 \<open>\<nu> \<star> \<pi>\<close> \<rho>
using \<nu>\<pi>.composable \<pi>\<rho>.composable by (unfold_locales, auto)
(* The two other cases, \<mu>\<nu>\<pi> and \<nu>\<pi>\<rho>, are part of the locale assumptions. *)
interpretation \<mu>_\<nu>\<pi>_\<rho>: three_composable_identity_arrows_of_spans C prj0 prj1 \<mu> \<open>\<nu> \<star> \<pi>\<close> \<rho> ..
interpretation \<mu>_\<nu>_\<pi>\<rho>: three_composable_identity_arrows_of_spans C prj0 prj1 \<mu> \<nu> \<open>\<pi> \<star> \<rho>\<close> ..
interpretation \<mu>\<nu>_\<pi>_\<rho>: three_composable_identity_arrows_of_spans C prj0 prj1 \<open>\<mu> \<star> \<nu>\<close> \<pi> \<rho> ..
lemma chines_eq:
shows "H\<mu>HH\<nu>\<pi>\<rho>.chine = \<mu>.leg0 \<down>\<down> HH\<nu>\<pi>\<rho>.leg1"
and "HH\<mu>H\<nu>\<pi>\<rho>.chine = assoc\<mu>\<nu>\<pi>.cod.leg0 \<down>\<down> \<rho>.leg1"
and "H\<mu>H\<nu>H\<pi>\<rho>.chine = \<mu>.leg0 \<down>\<down> H\<nu>H\<pi>\<rho>.leg1"
proof -
show "H\<mu>HH\<nu>\<pi>\<rho>.chine = \<mu>.leg0 \<down>\<down> HH\<nu>\<pi>\<rho>.leg1"
using hcomp_def [of \<mu> "hcomp (hcomp \<nu> \<pi>) \<rho>"] chine_hcomp_ide_ide \<mu>\<nu>.composable
by simp
show "HH\<mu>H\<nu>\<pi>\<rho>.chine = assoc\<mu>\<nu>\<pi>.cod.leg0 \<down>\<down> \<rho>.leg1"
by (simp add: \<mu>_\<nu>\<pi>.leg0_composite \<mu>_\<nu>\<pi>_\<rho>.chine_composite(1))
show "H\<mu>H\<nu>H\<pi>\<rho>.chine = \<mu>.leg0 \<down>\<down> H\<nu>H\<pi>\<rho>.leg1"
using hcomp_def [of \<mu> "hcomp \<nu> (hcomp \<pi> \<rho>)"] chine_hcomp_ide_ide \<mu>\<nu>.composable
by simp
qed
lemma cospan_\<mu>0_H\<nu>H\<pi>\<rho>1:
shows "C.cospan \<mu>.leg0 H\<nu>H\<pi>\<rho>.leg1"
by (metis C.cod_comp H\<nu>H\<pi>\<rho>.dom.leg_simps(3) \<mu>_\<nu>_\<pi>\<rho>.cospan_\<mu>\<nu> \<nu>\<pi>\<rho>.leg1_composite(1))
(* TODO: Better name for this. *)
lemma assoc_in_homs:
shows "\<guillemotleft>\<mu> \<star> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) : \<mu> \<star> (\<nu> \<star> \<pi>) \<star> \<rho> \<Rightarrow> \<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<guillemotright>"
and "\<guillemotleft>assoc\<^sub>S\<^sub>B \<mu> (\<nu> \<star> \<pi>) \<rho> : (\<mu> \<star> \<nu> \<star> \<pi>) \<star> \<rho> \<Rightarrow> \<mu> \<star> (\<nu> \<star> \<pi>) \<star> \<rho>\<guillemotright>"
and "\<guillemotleft>assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi> \<star> \<rho> : ((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho> \<Rightarrow> (\<mu> \<star> \<nu> \<star> \<pi>) \<star> \<rho>\<guillemotright>"
and "\<guillemotleft>assoc\<^sub>S\<^sub>B \<mu> \<nu> (\<pi> \<star> \<rho>) : (\<mu> \<star> \<nu>) \<star> \<pi> \<star> \<rho> \<Rightarrow> \<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<guillemotright>"
and "\<guillemotleft>assoc\<^sub>S\<^sub>B (\<mu> \<star> \<nu>) \<pi> \<rho> : ((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho> \<Rightarrow> (\<mu> \<star> \<nu>) \<star> \<pi> \<star> \<rho>\<guillemotright>"
proof -
show "\<guillemotleft>\<mu> \<star> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) : \<mu> \<star> (\<nu> \<star> \<pi>) \<star> \<rho> \<Rightarrow> \<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable by auto
show "\<guillemotleft>assoc\<^sub>S\<^sub>B \<mu> (\<nu> \<star> \<pi>) \<rho> : (\<mu> \<star> \<nu> \<star> \<pi>) \<star> \<rho> \<Rightarrow> \<mu> \<star> (\<nu> \<star> \<pi>) \<star> \<rho>\<guillemotright>"
using assoc_in_hom \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable by simp
show "\<guillemotleft>assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi> \<star> \<rho> : ((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho> \<Rightarrow> (\<mu> \<star> \<nu> \<star> \<pi>) \<star> \<rho>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable by auto
show "\<guillemotleft>assoc\<^sub>S\<^sub>B \<mu> \<nu> (\<pi> \<star> \<rho>) : (\<mu> \<star> \<nu>) \<star> \<pi> \<star> \<rho> \<Rightarrow> \<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable
by auto
show "\<guillemotleft>assoc\<^sub>S\<^sub>B (\<mu> \<star> \<nu>) \<pi> \<rho> : ((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho> \<Rightarrow> (\<mu> \<star> \<nu>) \<star> \<pi> \<star> \<rho>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable by auto
qed
lemma chine_composites:
shows "Chn (\<mu> \<star> assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) = chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)"
and "Chn (assoc\<^sub>S\<^sub>B \<mu> (\<nu> \<star> \<pi>) \<rho>) = \<mu>_\<nu>\<pi>_\<rho>.chine_assoc"
and "Chn (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi> \<star> \<rho>) = chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
and "Chn (assoc\<^sub>S\<^sub>B \<mu> \<nu> (\<pi> \<star> \<rho>)) = \<mu>_\<nu>_\<pi>\<rho>.chine_assoc"
and "Chn (assoc\<^sub>S\<^sub>B (\<mu> \<star> \<nu>) \<pi> \<rho>) = \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
proof -
show "Chn (\<mu> \<star> assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) = chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)"
using hcomp_def [of \<mu> "assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>"] chine_hcomp_def [of \<mu> "assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>"]
\<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable
by auto
show "Chn (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi> \<star> \<rho>) = chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using assoc_in_homs(2-3) hcomp_def
by (metis arrI arrow_of_spans_data.select_convs(1) hseqE)
show "Chn (assoc\<^sub>S\<^sub>B \<mu> (\<nu> \<star> \<pi>) \<rho>) = \<mu>_\<nu>\<pi>_\<rho>.chine_assoc"
using hcomp_def
by (meson arrow_of_spans_data.select_convs(1))
show "Chn (assoc\<^sub>S\<^sub>B \<mu> \<nu> (\<pi> \<star> \<rho>)) = \<mu>_\<nu>_\<pi>\<rho>.chine_assoc"
using hcomp_def
by (meson arrow_of_spans_data.select_convs(1))
show "Chn (assoc\<^sub>S\<^sub>B (\<mu> \<star> \<nu>) \<pi> \<rho>) = \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
using hcomp_def
by (meson arrow_of_spans_data.select_convs(1))
qed
lemma prj_in_homs [intro, simp]:
shows "\<guillemotleft>\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] : H\<mu>HH\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<nu>\<pi>\<rho>.chine\<guillemotright>"
and "\<guillemotleft>\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] : H\<mu>H\<nu>H\<pi>\<rho>.chine \<rightarrow>\<^sub>C \<mu>.apex\<guillemotright>"
and "\<guillemotleft>\<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.cod.leg0, \<rho>.cod.leg1] : HH\<mu>H\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>\<pi>.chine\<guillemotright>"
and "\<guillemotleft>\<p>\<^sub>0[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C \<rho>.chine\<guillemotright>"
and "\<guillemotleft>\<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>\<nu>\<pi>.chine\<guillemotright>"
and "\<guillemotleft>\<p>\<^sub>1[\<nu>\<pi>.leg0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0, \<rho>.leg1] : HH\<mu>H\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>\<pi>.chine\<guillemotright>"
and "\<guillemotleft>\<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>\<nu>\<pi>.chine\<guillemotright>"
and "\<guillemotleft>\<mu>_\<nu>\<pi>.prj\<^sub>0 : H\<mu>H\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<nu>\<pi>.apex\<guillemotright>"
proof -
show "\<guillemotleft>\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] : H\<mu>HH\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<nu>\<pi>\<rho>.chine\<guillemotright>"
by (metis C.dom_comp C.prj0_in_hom' C.prj1_simps_arr HH\<nu>\<pi>\<rho>.chine_eq_apex
HH\<nu>\<pi>\<rho>.cod.apex_def HH\<nu>\<pi>\<rho>.cod.is_span HH\<nu>\<pi>\<rho>.cod_simps(1) HH\<nu>\<pi>\<rho>.cod_simps(3)
\<mu>_\<nu>\<pi>_\<rho>.prj_simps(10) \<mu>_\<nu>\<pi>_\<rho>.prj_simps(11) \<mu>_\<nu>\<pi>_\<rho>.prj_simps(14) \<nu>\<pi>_\<rho>.leg1_composite)
show "\<guillemotleft>\<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.cod.leg0, \<rho>.cod.leg1] : HH\<mu>H\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>\<pi>.chine\<guillemotright>"
using \<mu>\<nu>\<pi>.leg0_composite(1) \<mu>\<nu>_\<pi>_\<rho>.prj_simps(3) \<mu>_\<nu>\<pi>.leg0_composite
\<mu>_\<nu>\<pi>_\<rho>.prj_simps(6)
by auto
show "\<guillemotleft>\<p>\<^sub>0[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C \<rho>.chine\<guillemotright>"
by (simp add: \<mu>\<nu>_\<pi>.leg0_composite \<mu>\<nu>_\<pi>_\<rho>.prj_in_hom(3))
show "\<guillemotleft>\<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>\<nu>\<pi>.chine\<guillemotright>"
using \<mu>\<nu>.leg0_composite \<mu>\<nu>_\<pi>.leg0_composite \<mu>\<nu>_\<pi>_\<rho>.prj_in_hom(2) by fastforce
show "\<guillemotleft>\<p>\<^sub>1[\<nu>\<pi>.leg0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0, \<rho>.leg1] : HH\<mu>H\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>\<pi>.chine\<guillemotright>"
using \<mu>\<nu>\<pi>.leg0_composite(1) \<mu>_\<nu>\<pi>.leg0_composite \<mu>_\<nu>\<pi>_\<rho>.prj_simps(3)
\<mu>_\<nu>\<pi>_\<rho>.prj_simps(6)
by force
show "\<guillemotleft>\<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>\<nu>\<pi>.chine\<guillemotright>"
using \<open>\<guillemotleft>\<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1] : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>\<nu>\<pi>.chine\<guillemotright>\<close> by fastforce
show "\<guillemotleft>\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] : H\<mu>H\<nu>H\<pi>\<rho>.chine \<rightarrow>\<^sub>C \<mu>.apex\<guillemotright>"
by (simp add: \<mu>_\<nu>_\<pi>\<rho>.prj_in_hom(4) \<nu>_\<pi>\<rho>.leg1_composite)
show "\<guillemotleft>\<mu>_\<nu>\<pi>.prj\<^sub>0 : H\<mu>H\<nu>\<pi>.chine \<rightarrow>\<^sub>C \<nu>\<pi>.apex\<guillemotright>"
using \<mu>_\<nu>\<pi>.chine_composite \<mu>_\<nu>\<pi>.prj_in_hom(2) by presburger
qed
lemma chine_in_homs [intro, simp]:
shows "\<guillemotleft>assoc\<mu>\<nu>\<pi>.chine : HH\<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>\<pi>.chine\<guillemotright>"
and "\<guillemotleft>assoc\<nu>\<pi>\<rho>.chine : HH\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
and "\<guillemotleft>chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) : H\<mu>HH\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
and "\<guillemotleft>chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>H\<nu>\<pi>\<rho>.chine\<guillemotright>"
proof -
show 1: "\<guillemotleft>assoc\<mu>\<nu>\<pi>.chine : HH\<mu>\<nu>\<pi>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>\<pi>.chine\<guillemotright>"
using \<mu>\<nu>\<pi>.chine_assoc_in_hom by simp
show "\<guillemotleft>assoc\<nu>\<pi>\<rho>.chine : HH\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
using \<nu>\<pi>\<rho>.chine_assoc_in_hom by simp
show "\<guillemotleft>chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) : H\<mu>HH\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<mu>H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
using Chn_in_hom by (metis assoc_in_homs(1) chine_composites(1))
show "\<guillemotleft>chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>H\<nu>\<pi>\<rho>.chine\<guillemotright>"
using Chn_in_hom by (metis assoc_in_homs(3) chine_composites(3))
qed
lemma commutative_squares [intro, simp]:
shows "C.commutative_square \<nu>\<pi>.leg0 \<rho>.leg1 \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0"
and "C.commutative_square \<nu>.leg0 \<pi>\<rho>.leg1 \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0"
and "C.commutative_square \<p>\<^sub>0[\<mu>.cod.leg0, assoc\<nu>\<pi>\<rho>.cod.leg1] assoc\<nu>\<pi>\<rho>.chine
(chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)) \<p>\<^sub>0[\<mu>.leg0, assoc\<nu>\<pi>\<rho>.dom.leg1]"
and "C.commutative_square \<p>\<^sub>1[\<mu>.cod.leg0, assoc\<nu>\<pi>\<rho>.cod.leg1] \<mu>.chine
(chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)) \<p>\<^sub>1[\<mu>.leg0, assoc\<nu>\<pi>\<rho>.dom.leg1]"
and "C.commutative_square assoc\<mu>\<nu>\<pi>.cod.leg0 \<rho>.cod.leg1
(assoc\<mu>\<nu>\<pi>.chine \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1])
(\<rho>.chine \<cdot> \<p>\<^sub>0[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1])"
and "C.commutative_square \<mu>.leg0 (\<nu>\<pi>.leg1 \<cdot> \<nu>\<pi>_\<rho>.prj\<^sub>1) \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1
\<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle>"
and "C.commutative_square \<mu>.leg0 (\<nu>.leg1 \<cdot> \<nu>_\<pi>\<rho>.prj\<^sub>1) \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>1\<^sub>1
\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle>"
proof -
show 1: "C.commutative_square \<nu>\<pi>.leg0 \<rho>.leg1 \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0"
proof -
have 1: "C.arr \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0 \<and> C.dom \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0 = HH\<mu>H\<nu>\<pi>\<rho>.chine \<and>
C.cod \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0 = \<rho>.apex"
by (meson C.in_homE \<mu>_\<nu>\<pi>_\<rho>.prj_in_hom(3))
hence "(\<nu>\<pi>.leg0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0) \<cdot> \<p>\<^sub>1[\<nu>\<pi>.leg0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0, \<rho>.leg1] = \<rho>.leg1 \<cdot> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0"
by (meson C.prj0_simps_arr C.pullback_commutes')
thus ?thesis
using 1 C.comp_assoc \<nu>\<pi>_\<rho>.legs_form_cospan(1) by simp
qed
show 2: "C.commutative_square \<nu>.leg0 \<pi>\<rho>.leg1 \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0"
proof -
have "\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0 \<cdot> \<p>\<^sub>1[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>\<rho>.leg1] =
\<pi>\<rho>.leg1 \<cdot> \<p>\<^sub>0[\<nu>.leg0 \<cdot> \<mu>\<nu>.prj\<^sub>0, \<pi>\<rho>.leg1]"
by (metis (no_types) C.category_axioms C.prj0_simps_arr C.pullback_commutes'
category.comp_reduce \<mu>_\<nu>_\<pi>\<rho>.prj_simps(2) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(3))
thus ?thesis
using C.commutative_square_def \<mu>_\<nu>_\<pi>\<rho>.cospan_\<nu>\<pi>
\<mu>_\<nu>_\<pi>\<rho>.prj_simps(2) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(3) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(5) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(6)
\<mu>_\<nu>_\<pi>\<rho>.prj_simps(8) \<nu>.dom.apex_def
by presburger
qed
show "C.commutative_square \<p>\<^sub>0[\<mu>.cod.leg0, assoc\<nu>\<pi>\<rho>.cod.leg1] assoc\<nu>\<pi>\<rho>.chine
(chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)) \<p>\<^sub>0[\<mu>.leg0, assoc\<nu>\<pi>\<rho>.dom.leg1]"
using assoc_in_homs(1) chine_hcomp_props(4) [of "assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>" \<mu>] hseq_char by blast
show "C.commutative_square \<p>\<^sub>1[\<mu>.cod.leg0, assoc\<nu>\<pi>\<rho>.cod.leg1] \<mu>.chine
(chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)) \<p>\<^sub>1[\<mu>.leg0, assoc\<nu>\<pi>\<rho>.dom.leg1]"
using chine_hcomp_props(3) [of "assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>" \<mu>] hseq_char
\<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable
by auto
show "C.commutative_square assoc\<mu>\<nu>\<pi>.cod.leg0 \<rho>.cod.leg1
(assoc\<mu>\<nu>\<pi>.chine \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1])
(\<rho>.chine \<cdot> \<p>\<^sub>0[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1])"
using assoc_in_homs(3) hseq_char chine_hcomp_props(2) by blast
show "C.commutative_square \<mu>.leg0 (\<nu>\<pi>.leg1 \<cdot> \<nu>\<pi>_\<rho>.prj\<^sub>1) \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1
\<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle>"
proof
show "C.cospan \<mu>.leg0 (\<nu>\<pi>.leg1 \<cdot> \<nu>\<pi>_\<rho>.prj\<^sub>1)"
using HH\<nu>\<pi>\<rho>.dom.leg_simps(1) \<mu>_\<nu>\<pi>_\<rho>.cospan_\<mu>\<nu> C.arrI by fastforce
show "C.span \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1 \<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle>"
using 1 \<mu>_\<nu>\<pi>_\<rho>.prj_in_hom(1) by auto
show "C.dom \<mu>.leg0 = C.cod \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1"
by simp
show "\<mu>.leg0 \<cdot> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1 =
(\<nu>\<pi>.leg1 \<cdot> \<nu>\<pi>_\<rho>.prj\<^sub>1) \<cdot> \<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle>"
by (metis (no_types, lifting) "1" C.comp_assoc C.prj_tuple(2) C.pullback_commutes'
\<mu>_\<nu>\<pi>_\<rho>.cospan_\<mu>\<nu>)
qed
show "C.commutative_square \<mu>.leg0 (\<nu>.leg1 \<cdot> \<nu>_\<pi>\<rho>.prj\<^sub>1) \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>1\<^sub>1
\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle>"
proof
show "C.cospan \<mu>.leg0 (\<nu>.leg1 \<cdot> \<nu>_\<pi>\<rho>.prj\<^sub>1)"
using C.arrI \<mu>_\<nu>_\<pi>\<rho>.prj_in_hom(4) by auto
show "C.span \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>1\<^sub>1 \<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle>"
using 2 by fastforce
thus "C.dom \<mu>.leg0 = C.cod \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>1\<^sub>1"
using \<mu>_\<nu>_\<pi>\<rho>.cospan_\<mu>\<nu> by simp
show "\<mu>.leg0 \<cdot> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>1\<^sub>1 =
(\<nu>.leg1 \<cdot> \<nu>_\<pi>\<rho>.prj\<^sub>1) \<cdot> \<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle>"
by (metis (no_types, lifting) "2" C.comp_assoc C.prj_tuple(2) C.pullback_commutes'
\<mu>_\<nu>_\<pi>\<rho>.cospan_\<mu>\<nu>)
qed
qed
lemma chine_pentagon:
shows "Chn ((\<mu> \<star> assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) \<bullet> assoc\<^sub>S\<^sub>B \<mu> (\<nu> \<star> \<pi>) \<rho> \<bullet> (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi> \<star> \<rho>)) =
Chn (assoc\<^sub>S\<^sub>B \<mu> \<nu> (\<pi> \<star> \<rho>) \<bullet> assoc\<^sub>S\<^sub>B (\<mu> \<star> \<nu>) \<pi> \<rho>)"
proof -
let ?LHS = "(\<mu> \<star> assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) \<bullet> assoc\<^sub>S\<^sub>B \<mu> (\<nu> \<star> \<pi>) \<rho> \<bullet> (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi> \<star> \<rho>)"
let ?RHS = "assoc\<^sub>S\<^sub>B \<mu> \<nu> (\<pi> \<star> \<rho>) \<bullet> assoc\<^sub>S\<^sub>B (\<mu> \<star> \<nu>) \<pi> \<rho>"
have LHS_in_hom: "\<guillemotleft>?LHS : ((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho> \<Rightarrow> \<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable by auto
have RHS_in_hom: "\<guillemotleft>?RHS : ((\<mu> \<star> \<nu>) \<star> \<pi>) \<star> \<rho> \<Rightarrow> \<mu> \<star> \<nu> \<star> \<pi> \<star> \<rho>\<guillemotright>"
using \<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable by auto
have 1: "arrow_of_spans (\<cdot>) ?LHS"
using arr_char assoc_in_homs(1-3) by blast
have L: "Chn ?LHS = chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using Chn_vcomp 1 arr_char chine_composites(1) chine_composites(3) seq_char
by fastforce
have R: "Chn ?RHS = \<mu>_\<nu>_\<pi>\<rho>.chine_assoc \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
using Chn_vcomp assoc_in_homs(4) assoc_in_homs(5) seqI' by auto
text \<open>
The outline of the proof is to show that the compositions of \<open>?LHS\<close>
and \<open>?RHS\<close> with the two projections \<open>\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1]\<close> and
\<open>\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1]\<close> are equal, and then apply \<open>\<nu>\<pi>\<rho>.prj'_joint_monic\<close>.
\<close>
text \<open>
The case for projection \<open>\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1]\<close> does not have subcases,
so we'll dispatch that one first.
\<close>
have "\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?LHS = \<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?RHS"
proof -
have "\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?LHS = \<mu>\<nu>.prj\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>_\<rho>.Prj\<^sub>1\<^sub>1"
proof -
have "\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?LHS =
\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using L by simp
also have "... = \<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
proof -
have "\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) =
\<mu>.chine \<cdot> \<p>\<^sub>1[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]"
proof -
have "C.commutative_square \<p>\<^sub>1[\<mu>.cod.leg0, assoc\<nu>\<pi>\<rho>.cod.leg1] \<mu>.chine
(chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)) \<p>\<^sub>1[\<mu>.leg0, assoc\<nu>\<pi>\<rho>.dom.leg1]"
by blast
thus ?thesis by auto
qed
thus ?thesis
using C.comp_permute [of "\<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1]" "chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)"
\<mu>.chine "\<p>\<^sub>1[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]"
"\<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"]
by blast
qed
also have "... = \<mu>.chine \<cdot> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using C.comp_reduce [of "\<p>\<^sub>1[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]" \<mu>_\<nu>\<pi>_\<rho>.chine_assoc]
\<nu>\<pi>_\<rho>.leg1_composite
by fastforce
also have "... = \<mu>.chine \<cdot> \<mu>\<nu>.prj\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>_\<rho>.Prj\<^sub>1\<^sub>1"
proof -
have "\<mu>.chine \<cdot> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>1\<^sub>1 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<mu>.chine \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>1 \<cdot> \<p>\<^sub>1[\<nu>\<pi>.leg0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0, \<rho>.leg1] \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using C.comp_assoc by simp
also have "... = \<mu>.chine \<cdot> (\<mu>_\<nu>\<pi>.prj\<^sub>1 \<cdot> assoc\<mu>\<nu>\<pi>.chine) \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1]"
proof -
have "\<p>\<^sub>1[\<nu>\<pi>.leg0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0, \<rho>.leg1] \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
assoc\<mu>\<nu>\<pi>.chine \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1]"
using chine_hcomp_props(6) [of \<rho> "assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>"] hcomp_def [of \<mu> "hcomp \<nu> \<pi>"]
\<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable
hseq_char assoc_in_homs(3)
by auto
thus ?thesis
using C.comp_assoc by auto
qed
also have "... = \<mu>.chine \<cdot> \<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1 \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1]"
using \<mu>\<nu>\<pi>.prj_chine_assoc(1) hcomp_def \<nu>\<pi>.composable by auto
also have "... = \<mu>.chine \<cdot> \<mu>\<nu>.prj\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>_\<rho>.Prj\<^sub>1\<^sub>1"
proof -
have "\<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>1 \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1] = \<mu>\<nu>.prj\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>_\<rho>.Prj\<^sub>1\<^sub>1"
by (simp add: C.comp_assoc \<mu>\<nu>.leg0_composite \<mu>\<nu>_\<pi>.leg0_composite)
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
also have "... = \<mu>\<nu>.prj\<^sub>1 \<cdot> \<mu>\<nu>_\<pi>_\<rho>.Prj\<^sub>1\<^sub>1"
using \<mu>\<nu>_\<pi>_\<rho>.prj_in_hom(1) hcomp_def [of \<mu> \<nu>] chine_hcomp_ide_ide \<mu>\<nu>.cod.apex_def
\<mu>\<nu>.composable \<mu>_\<nu>_\<pi>\<rho>.cospan_\<mu>\<nu> C.comp_ide_arr
by auto
finally show ?thesis by blast
qed
also have "... = \<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?RHS"
by (metis C.comp_assoc R \<mu>\<nu>.leg0_composite \<mu>\<nu>_\<pi>_\<rho>.prj_chine_assoc(1)
\<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(1) \<nu>_\<pi>\<rho>.leg1_composite \<pi>\<rho>.leg1_composite)
finally show ?thesis by blast
qed
text \<open>
Now for the case of \<open>\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1]\<close>.
We have to consider three sub-cases, involving the compositions with the projections
\<open>\<nu>\<pi>\<rho>.Prj\<^sub>1\<close>, \<open>\<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0\<close>, and \<open>\<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0\<close>.
\<close>
moreover have "\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?LHS =
\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?RHS"
proof -
(* Facts common to the three sub-cases. *)
have A: "\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc =
\<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle>"
using \<mu>_\<nu>\<pi>_\<rho>.chine_assoc_def \<nu>\<pi>_\<rho>.leg1_composite by auto
have B: "\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<mu>_\<nu>\<pi>.prj\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine_assoc \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
proof -
have "\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
(\<mu>_\<nu>\<pi>.prj\<^sub>0 \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.cod.leg0, \<rho>.cod.leg1]) \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using \<mu>_\<nu>\<pi>.composable \<nu>\<pi>.composite_is_arrow hcomp_def by auto
also have "... = \<mu>_\<nu>\<pi>.prj\<^sub>0 \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.cod.leg0, \<rho>.cod.leg1] \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using C.comp_assoc by simp
also have "... = \<mu>_\<nu>\<pi>.prj\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine_assoc \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
proof -
have "HH\<mu>\<nu>\<pi>.leg0 = assoc\<mu>\<nu>\<pi>.dom.leg0"
using hcomp_def [of "hcomp \<mu> \<nu>" \<pi>] by simp
moreover have "C.commutative_square assoc\<mu>\<nu>\<pi>.cod.leg0 \<rho>.cod.leg1
(assoc\<mu>\<nu>\<pi>.chine \<cdot> \<p>\<^sub>1[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1])
(\<rho>.chine \<cdot> \<p>\<^sub>0[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1])"
by blast
ultimately show ?thesis
using chine_hcomp_def [of "assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>" \<rho>] by simp
qed
finally show ?thesis by blast
qed
have *: "assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
proof -
text \<open>Subcase \<open>\<nu>\<pi>\<rho>.Prj\<^sub>1\<close>:\<close>
have "\<nu>\<pi>\<rho>.Prj\<^sub>1 \<cdot> assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<nu>\<pi>\<rho>.Prj\<^sub>1 \<cdot> \<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
proof -
have "\<nu>\<pi>\<rho>.Prj\<^sub>1 \<cdot> assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>1 \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using \<nu>\<pi>\<rho>.chine_assoc_props(1) C.prj0_in_hom [of \<mu>.leg0 HH\<nu>\<pi>\<rho>.leg1] cospan_\<mu>0_H\<nu>H\<pi>\<rho>1
C.comp_reduce [of \<nu>\<pi>\<rho>.Prj\<^sub>1 assoc\<nu>\<pi>\<rho>.chine \<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>1
"\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"]
by auto
also have "... = \<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>1 \<cdot> \<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle> \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using A C.comp_reduce [of "\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]" \<mu>_\<nu>\<pi>_\<rho>.chine_assoc]
by fastforce
also have "... = \<nu>\<pi>.prj\<^sub>1 \<cdot> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
by (metis A C.comp_assoc \<mu>_\<nu>\<pi>_\<rho>.prj_chine_assoc(2) \<nu>\<pi>.leg0_composite
\<nu>\<pi>_\<rho>.leg1_composite)
also have "... = \<nu>\<pi>.prj\<^sub>1 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine_assoc \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
using B by simp
also have "... = \<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>1 \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
using hcomp_def [of \<nu> \<pi>] \<nu>\<pi>.composable C.comp_assoc
C.comp_reduce [of \<mu>\<nu>\<pi>.Prj\<^sub>1\<^sub>0 \<mu>\<nu>\<pi>.chine_assoc \<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>1 "\<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"]
\<mu>\<nu>\<pi>.prj_in_hom(5) \<mu>\<nu>\<pi>.prj_chine_assoc(2)
by auto
also have "... = \<nu>\<pi>\<rho>.Prj\<^sub>1 \<cdot>
\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
by (metis (no_types, lifting) C.comp_assoc C.prj_tuple(2) \<mu>\<nu>.leg0_composite
\<mu>\<nu>\<pi>.leg0_composite(2) \<mu>\<nu>_\<pi>_\<rho>.prj_chine_assoc(1) \<pi>\<rho>.leg1_composite
commutative_squares(2))
finally show ?thesis by simp
qed
moreover
text \<open>Subcase \<open>\<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0\<close>:\<close>
have "\<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0 \<cdot> assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0 \<cdot> \<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
proof -
have "\<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0 \<cdot> assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using C.comp_reduce [of \<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0 "assoc\<nu>\<pi>\<rho>.chine" \<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>1
"\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"]
by auto
also have "... = \<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<cdot> \<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle> \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using A C.comp_reduce [of "\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]" \<mu>_\<nu>\<pi>_\<rho>.chine_assoc]
by fastforce
also have "... = \<nu>\<pi>.prj\<^sub>0 \<cdot> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
by (metis A C.comp_assoc \<mu>_\<nu>\<pi>_\<rho>.prj_chine_assoc(2) \<nu>\<pi>.leg0_composite
\<nu>\<pi>_\<rho>.leg1_composite)
also have "... = \<nu>\<pi>.prj\<^sub>0 \<cdot> \<mu>_\<nu>\<pi>.prj\<^sub>0 \<cdot> \<mu>\<nu>\<pi>.chine_assoc \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
using B by simp
also have "... = \<mu>\<nu>\<pi>.Prj\<^sub>0 \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
using hcomp_def [of \<nu> \<pi>] \<nu>\<pi>.composable \<mu>\<nu>\<pi>.prj_in_hom(6)
C.comp_reduce [of \<nu>\<pi>.prj\<^sub>0 \<mu>_\<nu>\<pi>.prj\<^sub>0 \<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0
"\<mu>\<nu>\<pi>.chine_assoc \<cdot> \<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"]
C.comp_reduce [of \<mu>\<nu>\<pi>.Prj\<^sub>0\<^sub>0 \<mu>\<nu>\<pi>.chine_assoc \<mu>\<nu>\<pi>.Prj\<^sub>0
"\<p>\<^sub>1[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"]
by fastforce
also have "... = \<nu>\<pi>\<rho>.Prj\<^sub>1\<^sub>0 \<cdot>
\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
by (metis C.comp_assoc C.tuple_prj \<mu>\<nu>.leg0_composite \<mu>\<nu>_\<pi>.leg0_composite
\<mu>\<nu>_\<pi>_\<rho>.prj_chine_assoc(2) \<mu>_\<nu>_\<pi>\<rho>.cospan_\<nu>\<pi> \<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(2)
\<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(3) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(2) \<pi>\<rho>.leg1_composite)
finally show ?thesis by blast
qed
moreover
text \<open>Subcase \<open>\<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0\<close>:\<close>
have "\<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0 \<cdot> assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0 \<cdot> \<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
proof -
have "\<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0 \<cdot> assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> =
\<nu>\<pi>\<rho>.Prj\<^sub>0 \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using C.comp_reduce [of \<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0 assoc\<nu>\<pi>\<rho>.chine \<nu>\<pi>\<rho>.Prj\<^sub>0
"\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"]
by fastforce
also have "... = \<nu>\<pi>\<rho>.Prj\<^sub>0 \<cdot> \<langle>\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>\<pi>.leg0, \<rho>.leg1\<rbrakk> \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0\<rangle> \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using A C.comp_reduce [of "\<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]" \<mu>_\<nu>\<pi>_\<rho>.chine_assoc]
by fastforce
also have "... = \<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0 \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
by (metis A C.comp_assoc \<mu>_\<nu>\<pi>_\<rho>.prj_chine_assoc(3) \<nu>\<pi>.leg0_composite
\<nu>\<pi>_\<rho>.leg1_composite)
also have "... = \<rho>.chine \<cdot> \<p>\<^sub>0[assoc\<mu>\<nu>\<pi>.dom.leg0, \<rho>.leg1]"
proof -
have "\<mu>_\<nu>\<pi>_\<rho>.Prj\<^sub>0 = \<p>\<^sub>0[assoc\<mu>\<nu>\<pi>.cod.leg0, \<rho>.cod.leg1]"
using \<mu>_\<nu>\<pi>.composable \<nu>\<pi>.composite_is_arrow hcomp_def by auto
thus ?thesis
using chine_hcomp_props(5) [of \<rho> "assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>"]
\<mu>\<nu>.composable \<nu>\<pi>.composable \<pi>\<rho>.composable
by simp
qed
also have "... = \<p>\<^sub>0[HH\<mu>\<nu>\<pi>.leg0, \<rho>.leg1]"
by (metis C.comp_cod_arr C.in_homE arrow_of_spans_data.select_convs(2)
prj_in_homs(4))
also have "... = \<nu>\<pi>\<rho>.Prj\<^sub>0\<^sub>0 \<cdot>
\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
by (metis C.comp_assoc C.tuple_prj \<mu>\<nu>.leg0_composite \<mu>\<nu>_\<pi>.leg0_composite
\<mu>\<nu>_\<pi>_\<rho>.prj_chine_assoc(3) \<mu>_\<nu>_\<pi>\<rho>.cospan_\<nu>\<pi> \<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(2)
\<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(3) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(2) \<pi>\<rho>.leg1_composite)
finally show ?thesis by blast
qed
moreover have "\<guillemotleft>assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho> : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
using \<nu>\<pi>\<rho>.chine_assoc_props(1) by fast
moreover have "\<guillemotleft>\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc :
HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
proof -
have "\<guillemotleft>\<mu>\<nu>_\<pi>_\<rho>.chine_assoc : HHH\<mu>\<nu>\<pi>\<rho>.chine \<rightarrow>\<^sub>C HH\<mu>\<nu>H\<pi>\<rho>.chine\<guillemotright>"
using \<mu>\<nu>_\<pi>_\<rho>.chine_assoc_props(1) by blast
moreover have "\<guillemotleft>\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> :
HH\<mu>\<nu>H\<pi>\<rho>.chine \<rightarrow>\<^sub>C H\<nu>H\<pi>\<rho>.chine\<guillemotright>"
using chine_hcomp_ide_ide hcomp_def [of \<nu> "hcomp \<pi> \<rho>"] \<nu>\<pi>.composable by auto
ultimately show ?thesis by blast
qed
ultimately show ?thesis
using \<nu>\<pi>\<rho>.prj'_joint_monic
[of "assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
HHH\<mu>\<nu>\<pi>\<rho>.chine
"\<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"]
by simp
qed
text \<open>
Now use fact \<open>*\<close> to finish off the \<open>\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1]\<close> case.
\<close>
have "\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?LHS =
assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
proof -
have "\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?LHS =
\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
using L by simp
also have "... = assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1] \<cdot> \<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot>
chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>"
proof -
have "\<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>) =
assoc\<nu>\<pi>\<rho>.chine \<cdot> \<p>\<^sub>0[\<mu>.leg0, HH\<nu>\<pi>\<rho>.leg1]"
proof -
have "C.commutative_square \<p>\<^sub>0[\<mu>.cod.leg0, assoc\<nu>\<pi>\<rho>.cod.leg1] assoc\<nu>\<pi>\<rho>.chine
(chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>)) \<p>\<^sub>0[\<mu>.leg0, assoc\<nu>\<pi>\<rho>.dom.leg1]"
by blast
thus ?thesis
using \<nu>\<pi>\<rho>.chine_assoc_props(1) by auto
qed
moreover have "C.seq \<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] (chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>))"
using cospan_\<mu>0_H\<nu>H\<pi>\<rho>1 prj_in_homs(2) by fastforce
moreover have "C.seq (chine_hcomp \<mu> (assoc\<^sub>S\<^sub>B \<nu> \<pi> \<rho>))
(\<mu>_\<nu>\<pi>_\<rho>.chine_assoc \<cdot> chine_hcomp (assoc\<^sub>S\<^sub>B \<mu> \<nu> \<pi>) \<rho>)"
by blast
ultimately show ?thesis
using chine_hcomp_props(4) C.comp_permute by auto
qed
finally show ?thesis by blast
qed
also have "... = \<langle>\<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<^sub>1 \<lbrakk>\<nu>.leg0, \<pi>\<rho>.leg1\<rbrakk> \<mu>_\<nu>_\<pi>\<rho>.Prj\<^sub>0\<rangle> \<cdot> \<mu>\<nu>_\<pi>_\<rho>.chine_assoc"
using * by simp
also have "... = \<p>\<^sub>0[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] \<cdot> Chn ?RHS"
by (metis C.comp_assoc C.tuple_prj R \<mu>_\<nu>_\<pi>\<rho>.cospan_\<nu>\<pi> \<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(2)
\<mu>_\<nu>_\<pi>\<rho>.prj_chine_assoc(3) \<mu>_\<nu>_\<pi>\<rho>.prj_simps(2) \<nu>_\<pi>\<rho>.leg1_composite)
finally show ?thesis by blast
qed
moreover have "C.seq \<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] (Chn ?LHS)"
using LHS_in_hom Chn_in_hom by blast
moreover have "C.seq \<p>\<^sub>1[\<mu>.leg0, H\<nu>H\<pi>\<rho>.leg1] (Chn ?RHS)"
using RHS_in_hom Chn_in_hom by blast
ultimately show "Chn ?LHS = Chn ?RHS"
using cospan_\<mu>0_H\<nu>H\<pi>\<rho>1 C.prj_joint_monic by blast
qed
end
context span_bicategory
begin
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 "(f \<star> \<alpha>\<^sub>S\<^sub>B (g, h, k)) \<bullet> \<alpha>\<^sub>S\<^sub>B (f, g \<star> h, k) \<bullet> (\<alpha>\<^sub>S\<^sub>B (f, g, h) \<star> k) =
\<alpha>\<^sub>S\<^sub>B (f, g, h \<star> k) \<bullet> \<alpha>\<^sub>S\<^sub>B (f \<star> g, h, k)"
proof -
interpret f: identity_arrow_of_spans C f
using assms ide_char' by auto
interpret g: identity_arrow_of_spans C g
using assms ide_char' by auto
interpret h: identity_arrow_of_spans C h
using assms ide_char' by auto
interpret k: identity_arrow_of_spans C k
using assms ide_char' by auto
interpret fghk: four_composable_identity_arrows_of_spans C prj0 prj1 f g h k
using assms by (unfold_locales, auto)
let ?LHS = "(f \<star> assoc\<^sub>S\<^sub>B g h k) \<bullet> (assoc\<^sub>S\<^sub>B f (g \<star> h) k) \<bullet> (assoc\<^sub>S\<^sub>B f g h \<star> k)"
let ?RHS = "assoc\<^sub>S\<^sub>B f g (h \<star> k) \<bullet> assoc\<^sub>S\<^sub>B (f \<star> g) h k"
have "(f \<star> \<alpha>\<^sub>S\<^sub>B (g, h, k)) \<bullet> \<alpha>\<^sub>S\<^sub>B (f, g \<star> h, k) \<bullet> (\<alpha>\<^sub>S\<^sub>B (f, g, h) \<star> k) = ?LHS"
using assms \<alpha>_ide ide_hcomp src_hcomp trg_hcomp by simp
also have "... = ?RHS"
using fghk.\<mu>\<nu>.composable fghk.\<nu>\<pi>.composable fghk.\<pi>\<rho>.composable fghk.chine_pentagon
by (intro arr_eqI, auto)
also have "... = \<alpha>\<^sub>S\<^sub>B (f, g, h \<star> k) \<bullet> \<alpha>\<^sub>S\<^sub>B (f \<star> g, h, k)"
using assms \<alpha>_ide ide_hcomp src_hcomp trg_hcomp by simp
finally show ?thesis by blast
qed
lemma extends_to_bicategory:
shows "bicategory vcomp hcomp assoc unit src trg"
using unit_in_hom obj_char iso_unit assoc_def pentagon
apply unfold_locales by auto
sublocale bicategory vcomp hcomp assoc unit src trg
using extends_to_bicategory by auto
subsection "Miscellaneous Formulas"
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation in_hom ("\<guillemotleft>_ : _ \<Rightarrow> _\<guillemotright>")
notation lunit ("\<l>[_]")
notation runit ("\<r>[_]")
notation lunit' ("\<l>\<^sup>-\<^sup>1[_]")
notation runit' ("\<r>\<^sup>-\<^sup>1[_]")
notation assoc ("\<a>[_, _, _]")
notation \<a>' ("\<a>\<^sup>-\<^sup>1[_, _, _]")
lemma \<alpha>'_ide:
assumes "ide f" and "ide g" and "ide h"
and "src f = trg g" and "src g = trg h"
shows "\<alpha>' (f, g, h) = assoc'\<^sub>S\<^sub>B f g h"
proof -
have fgh: "VVV.ide (f, g, h)"
- using assms VVV.ide_char VVV.arr_char VV.arr_char by simp
+ using assms VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
interpret f: arrow_of_spans C f
using assms arr_char [of f] by auto
interpret g: arrow_of_spans C g
using assms arr_char [of g] by auto
interpret h: arrow_of_spans C h
using assms arr_char [of h] by auto
interpret fgh: three_composable_arrows_of_spans C prj0 prj1 f g h
using assms by (unfold_locales, auto)
interpret fgh: three_composable_identity_arrows_of_spans C prj0 prj1 f g h
using assms ide_char by (unfold_locales, auto)
have "\<alpha>' (f, g, h) = inv (\<alpha> (f, g, h))"
using fgh \<alpha>'.inverts_components
by (simp add: \<alpha>_def)
moreover have "inv (\<alpha> (f, g, h)) = \<lparr>Chn = C.inv (Chn (\<alpha> (f, g, h))),
Dom = Cod (\<alpha> (f, g, h)),
Cod = Dom (\<alpha> (f, g, h))\<rparr>"
using fgh \<alpha>.components_are_iso inv_eq
by (simp add: \<alpha>_def fgh.\<mu>\<nu>.composable fgh.\<nu>\<pi>.composable)
moreover have "... = assoc'\<^sub>S\<^sub>B f g h"
using assms fgh \<alpha>_ide [of f g h] fgh.chine_assoc_inverse C.inverse_unique
by (simp add: \<alpha>_def)
ultimately show ?thesis by simp
qed
text \<open>
The following give explicit expressions for the unitors,
derived from their characterizing properties and the definition of the associators.
\<close>
lemma runit_ide_eq:
assumes "ide f"
shows "\<r>[f] = \<lparr>Chn = \<p>\<^sub>1[Leg0 (Dom f), C.cod (Leg0 (Dom f))],
Dom = \<lparr>Leg0 = \<p>\<^sub>0[Leg0 (Dom f), C.cod (Leg0 (Dom f))],
Leg1 = Leg1 (Dom f) \<cdot> \<p>\<^sub>1[Leg0 (Dom f), C.cod (Leg0 (Dom f))]\<rparr>,
Cod = Cod f\<rparr>"
proof -
interpret f: identity_arrow_of_spans C f
using assms ide_char' by auto
interpret src: identity_arrow_of_spans C \<open>src f\<close>
using assms ide_char' ide_src by auto
interpret f_src: two_composable_identity_arrows_of_spans C prj0 prj1 f \<open>src f\<close>
using assms by (unfold_locales, simp)
interpret src_src: two_composable_identity_arrows_of_spans C prj0 prj1 \<open>src f\<close> \<open>src f\<close>
by (unfold_locales, simp)
interpret f_src_src: three_composable_identity_arrows_of_spans C prj0 prj1 f \<open>src f\<close> \<open>src f\<close>
..
let ?rf = "\<lparr>Chn = \<p>\<^sub>1[f.leg0, f.dsrc],
Dom = \<lparr>Leg0 = \<p>\<^sub>0[f.leg0, f.dsrc], Leg1 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.dsrc]\<rparr>,
Cod = Cod f\<rparr>"
have "?rf = \<r>[f]"
proof (intro runit_eqI)
show "ide f" by fact
interpret rf: arrow_of_spans C ?rf
proof -
interpret dom_rf: span_in_category C
\<open>\<lparr>Leg0 = \<p>\<^sub>0[f.leg0, f.dsrc], Leg1 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.dsrc]\<rparr>\<close>
by (unfold_locales, simp_all)
show "arrow_of_spans C ?rf"
using dom_rf.apex_def C.comp_cod_arr C.pullback_commutes [of f.leg0 f.dsrc]
apply unfold_locales by auto
qed
show rf_in_hom: "\<guillemotleft>?rf : f \<star> src f \<Rightarrow> f\<guillemotright>"
proof
show "arr ?rf"
using rf.arrow_of_spans_axioms arr_char by simp
show "cod ?rf = f"
using cod_char rf.arrow_of_spans_axioms arr_char by simp
show "dom ?rf = f \<star> src f"
using dom_char rf.arrow_of_spans_axioms src.arrow_of_spans_axioms arr_char hcomp_def
f.arrow_of_spans_axioms f_src.composable chine_hcomp_ide_ide src_def ide_char
C.comp_cod_arr rf.dom.apex_def
by simp
qed
show "?rf \<star> src f = (f \<star> \<i>[src f]) \<bullet> \<a>[f, src f, src f]"
proof (intro arr_eqI)
show par: "par (?rf \<star> src f) ((f \<star> \<i>[src f]) \<bullet> \<a>[f, src f, src f])"
proof -
have "\<guillemotleft>?rf \<star> src f : (f \<star> src f) \<star> src f \<Rightarrow> f \<star> src f\<guillemotright>"
using f_src_src.composites_are_arrows(2) rf_in_hom src_src.are_identities(1)
by blast
thus ?thesis by auto
qed
show "Chn (?rf \<star> src f) = Chn ((f \<star> \<i>[src f]) \<bullet> \<a>[f, src f, src f])"
proof -
have "Chn (?rf \<star> src f) = \<langle>f_src_src.Prj\<^sub>1\<^sub>1 \<lbrakk>f.leg0, src.leg1\<rbrakk> f_src_src.Prj\<^sub>0\<^sub>1\<rangle>"
proof -
have "Chn (?rf \<star> src f) =
\<langle>f_src_src.Prj\<^sub>1\<^sub>1 \<lbrakk>f.leg0, src.leg1\<rbrakk> \<p>\<^sub>0[f_src.prj\<^sub>0, src.leg1]\<rangle>"
using assms src_def trg_def hcomp_def arr_char ide_char
rf.arrow_of_spans_axioms src.identity_arrow_of_spans_axioms
chine_hcomp_arr_ide C.comp_cod_arr
by (simp add: f.arrow_of_spans_axioms identity_arrow_of_spans_def)
moreover have "\<p>\<^sub>0[f_src.prj\<^sub>0, src.leg1] = f_src_src.Prj\<^sub>0\<^sub>1"
proof -
have "src f = \<lparr>Chn = f.dsrc,
Dom = \<lparr>Leg0 = f.dsrc, Leg1 = f.dsrc\<rparr>,
Cod = \<lparr>Leg0 = f.dsrc, Leg1 = f.dsrc\<rparr>\<rparr>"
using assms src_def by simp
thus ?thesis
by (simp add: C.comp_cod_arr C.pullback_commutes')
qed
ultimately show ?thesis by auto
qed
also have "... = Chn ((f \<star> \<i>[src f]) \<bullet> \<a>[f, src f, src f])"
proof -
have "Chn ((f \<star> \<i>[src f]) \<bullet> \<a>[f, src f, src f]) =
\<langle>f_src_src.Prj\<^sub>1 \<lbrakk>f.leg0, src.leg1\<rbrakk> f_src_src.Prj\<^sub>1\<^sub>0\<rangle> \<cdot> f_src_src.chine_assoc"
proof -
have "Chn ((f \<star> \<i>[src f]) \<bullet> \<a>[f, src f, src f]) =
Chn (f \<star> \<i>[src f]) \<cdot> Chn \<a>[f, src f, src f]"
using par vcomp_eq [of "f \<star> \<i>[src f]" "\<a>[f, src f, src f]"]
by simp
moreover have "Chn (f \<star> \<i>[src f]) =
\<langle>f_src_src.Prj\<^sub>1 \<lbrakk>f.leg0, src.leg1\<rbrakk> f_src_src.Prj\<^sub>1\<^sub>0\<rangle>"
proof -
have "\<i>[src f] = \<lparr>Chn = \<p>\<^sub>1[f.dsrc, f.dsrc],
Dom = \<lparr>Leg0 = \<p>\<^sub>1[f.dsrc, f.dsrc], Leg1 = \<p>\<^sub>1[f.dsrc, f.dsrc]\<rparr>,
Cod = \<lparr>Leg0 = f.dsrc, Leg1 = f.dsrc\<rparr>\<rparr>"
using unit_def src_def trg_def hcomp_def src.arrow_of_spans_axioms arr_char
f.arrow_of_spans_axioms C.comp_cod_arr
by simp
moreover have "arrow_of_spans C \<i>[src f]"
using assms arr_char [of "\<i>[src f]"] by simp
ultimately show ?thesis
using assms unit_def hcomp_def chine_hcomp_ide_arr
rf.arrow_of_spans_axioms src.arrow_of_spans_axioms
f.arrow_of_spans_axioms arr_char C.comp_cod_arr
src_def trg_def
by simp
qed
ultimately show ?thesis
using \<alpha>_ide by simp
qed
also have "... = \<langle>f_src_src.Prj\<^sub>1 \<cdot> f_src_src.chine_assoc
\<lbrakk>f.leg0, src.leg1\<rbrakk>
f_src_src.Prj\<^sub>1\<^sub>0 \<cdot> f_src_src.chine_assoc\<rangle>"
using C.comp_assoc C.comp_tuple_arr C.pullback_commutes'
f_src_src.cospan_\<mu>\<nu> f_src_src.cospan_\<nu>\<pi>
by simp
also have "... = \<langle>f_src_src.Prj\<^sub>1\<^sub>1 \<lbrakk>f.leg0, src.leg1\<rbrakk> f_src_src.Prj\<^sub>0\<^sub>1\<rangle>"
by simp
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
qed
qed
thus ?thesis by simp
qed
lemma lunit_ide_eq:
assumes "ide f"
shows "\<l>[f] = \<lparr>Chn = \<p>\<^sub>0[C.cod (Leg1 (Dom f)), Leg1 (Dom f)],
Dom = \<lparr>Leg0 = Leg0 (Dom f) \<cdot> \<p>\<^sub>0[C.cod (Leg1 (Dom f)), Leg1 (Dom f)],
Leg1 = \<p>\<^sub>1[C.cod (Leg1 (Dom f)), Leg1 (Dom f)]\<rparr>,
Cod = Cod f\<rparr>"
proof -
interpret f: identity_arrow_of_spans C f
using assms ide_char' by auto
interpret trg: identity_arrow_of_spans C \<open>trg f\<close>
using assms ide_char' ide_trg by auto
interpret trg_f: two_composable_identity_arrows_of_spans C prj0 prj1 \<open>trg f\<close> f
using assms by (unfold_locales, simp)
interpret trg_trg: two_composable_identity_arrows_of_spans C prj0 prj1 \<open>trg f\<close> \<open>trg f\<close>
by (unfold_locales, simp)
interpret trg_trg_f: three_composable_identity_arrows_of_spans C prj0 prj1 \<open>trg f\<close> \<open>trg f\<close> f
..
let ?lf = "\<lparr>Chn = \<p>\<^sub>0[f.dtrg, f.leg1],
Dom = \<lparr>Leg0 = f.leg0 \<cdot> \<p>\<^sub>0[f.dtrg, f.leg1], Leg1 = \<p>\<^sub>1[f.dtrg, f.leg1]\<rparr>,
Cod = Cod f\<rparr>"
have "?lf = \<l>[f]"
proof (intro lunit_eqI)
show "ide f" by fact
interpret lf: arrow_of_spans C ?lf
proof -
interpret dom_lf: span_in_category C
\<open>\<lparr>Leg0 = f.leg0 \<cdot> \<p>\<^sub>0[f.dtrg, f.leg1], Leg1 = \<p>\<^sub>1[f.dtrg, f.leg1]\<rparr>\<close>
by (unfold_locales, simp_all)
show "arrow_of_spans C ?lf"
using dom_lf.apex_def C.comp_cod_arr C.pullback_commutes [of f.dtrg f.leg1]
apply unfold_locales by auto
qed
show lf_in_hom: "\<guillemotleft>?lf : trg f \<star> f \<Rightarrow> f\<guillemotright>"
proof
show "arr ?lf"
using lf.arrow_of_spans_axioms arr_char by simp
show "cod ?lf = f"
using cod_char lf.arrow_of_spans_axioms arr_char by simp
show "dom ?lf = trg f \<star> f"
using dom_char lf.arrow_of_spans_axioms trg.arrow_of_spans_axioms arr_char hcomp_def
f.arrow_of_spans_axioms trg_f.composable chine_hcomp_ide_ide trg_def ide_char
C.comp_cod_arr lf.dom.apex_def
by simp
qed
show "trg f \<star> ?lf = (\<i>[trg f] \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]"
proof (intro arr_eqI)
show par: "par (trg f \<star> ?lf) ((\<i>[trg f] \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[trg f, trg f, f])"
proof -
have "\<guillemotleft>trg f \<star> ?lf : trg f \<star> (trg f \<star> f) \<Rightarrow> trg f \<star> f\<guillemotright>"
proof -
have "trg f \<star> ?lf = L ?lf"
using assms lf_in_hom src_def trg_def arr_char lf.arrow_of_spans_axioms
f.arrow_of_spans_axioms
by simp
moreover have "\<guillemotleft>L ?lf : trg f \<star> (trg f \<star> f) \<Rightarrow> trg f \<star> f\<guillemotright>"
using lf_in_hom L.preserves_hom [of ?lf "trg f \<star> f" f] by simp
ultimately show ?thesis by auto
qed
thus ?thesis by auto
qed
show "Chn (trg f \<star> ?lf) = Chn ((\<i>[trg f] \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[trg f, trg f, f])"
proof -
have "Chn (trg f \<star> ?lf) = \<langle>trg_trg_f.Prj\<^sub>1\<^sub>0 \<lbrakk>trg.leg0, f.leg1\<rbrakk> trg_trg_f.Prj\<^sub>0\<^sub>0\<rangle>"
proof -
have "Chn (trg f \<star> ?lf) =
\<langle>\<p>\<^sub>1[trg.leg0, trg_f.prj\<^sub>1] \<lbrakk>trg.leg0, f.leg1\<rbrakk> trg_trg_f.Prj\<^sub>0\<^sub>0\<rangle>"
using assms src_def trg_def hcomp_def arr_char ide_char
lf.arrow_of_spans_axioms trg.identity_arrow_of_spans_axioms
chine_hcomp_ide_arr C.comp_cod_arr
by (simp add: f.arrow_of_spans_axioms identity_arrow_of_spans_def)
moreover have "\<p>\<^sub>1[trg.leg0, trg_f.prj\<^sub>1] = trg_trg_f.Prj\<^sub>1\<^sub>0"
proof -
have "trg f = \<lparr>Chn = f.dtrg,
Dom = \<lparr>Leg0 = f.dtrg, Leg1 = f.dtrg\<rparr>,
Cod = \<lparr>Leg0 = f.dtrg, Leg1 = f.dtrg\<rparr>\<rparr>"
using assms trg_def by simp
thus ?thesis
apply (simp add: C.comp_cod_arr C.pullback_commutes')
by (metis C.comp_cod_arr C.pullback_commutes' select_convs(1) select_convs(2)
select_convs(3) f.cod_simps(3) lf.cod_trg_eq_dom_trg lf.dom.leg_simps(3)
span_data.select_convs(1) span_data.select_convs(2) trg.chine_eq_apex
trg_trg_f.cospan_\<nu>\<pi> trg_trg_f.prj_simps(10) trg_trg_f.prj_simps(16))
qed
ultimately show ?thesis by auto
qed
also have "... = Chn ((\<i>[trg f] \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[trg f, trg f, f])"
proof -
have "Chn ((\<i>[trg f] \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]) =
\<langle>trg_trg_f.Prj\<^sub>0\<^sub>1 \<lbrakk>trg.leg0, f.leg1\<rbrakk> trg_trg_f.Prj\<^sub>0\<rangle> \<cdot> trg_trg_f.chine_assoc'"
proof -
have "Chn ((\<i>[trg f] \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[trg f, trg f, f]) =
Chn (\<i>[trg f] \<star> f) \<cdot> Chn \<a>\<^sup>-\<^sup>1[trg f, trg f, f]"
using par vcomp_eq [of "\<i>[trg f] \<star> f" "\<a>\<^sup>-\<^sup>1[trg f, trg f, f]"]
by simp
moreover have "Chn (\<i>[trg f] \<star> f) =
\<langle>trg_trg_f.Prj\<^sub>0\<^sub>1 \<lbrakk>trg.leg0, f.leg1\<rbrakk> trg_trg_f.Prj\<^sub>0\<rangle>"
proof -
have "\<i>[trg f] = \<lparr>Chn = \<p>\<^sub>1[f.dtrg, f.dtrg],
Dom = \<lparr>Leg0 = \<p>\<^sub>1[f.dtrg, f.dtrg], Leg1 = \<p>\<^sub>1[f.dtrg, f.dtrg]\<rparr>,
Cod = \<lparr>Leg0 = f.dtrg, Leg1 = f.dtrg\<rparr>\<rparr>"
using unit_def src_def trg_def hcomp_def trg.arrow_of_spans_axioms arr_char
f.arrow_of_spans_axioms C.comp_cod_arr
by simp
moreover have "arrow_of_spans C \<i>[trg f]"
using assms arr_char [of "\<i>[trg f]"] by simp
ultimately show ?thesis
using assms unit_def hcomp_def chine_hcomp_arr_ide
lf.arrow_of_spans_axioms trg.arrow_of_spans_axioms
f.arrow_of_spans_axioms arr_char C.comp_cod_arr
src_def trg_def
by simp
qed
moreover have "Chn \<a>\<^sup>-\<^sup>1[trg f, trg f, f] = trg_trg_f.chine_assoc'"
proof -
have "iso (\<alpha> (trg f, trg f, f))"
by (simp add: \<alpha>_def)
moreover have "C.inv trg_trg_f.chine_assoc = trg_trg_f.chine_assoc'"
using trg_trg_f.chine_assoc_inverse C.inv_is_inverse C.inverse_arrow_unique
by auto
ultimately show ?thesis
using assms by (simp add: \<a>'_def \<alpha>'_ide)
qed
ultimately show ?thesis
by simp
qed
also have "... = \<langle>trg_trg_f.Prj\<^sub>0\<^sub>1 \<cdot> trg_trg_f.chine_assoc'
\<lbrakk>trg.leg0, f.leg1\<rbrakk>
trg_trg_f.Prj\<^sub>0 \<cdot> trg_trg_f.chine_assoc'\<rangle>"
proof -
have "C.commutative_square trg.leg0 f.leg1 trg_trg_f.Prj\<^sub>0\<^sub>1 trg_trg_f.Prj\<^sub>0"
proof
show "C.cospan trg.leg0 f.leg1"
using trg_f.legs_form_cospan(1) by auto
show "C.span trg_trg_f.Prj\<^sub>0\<^sub>1 trg_trg_f.Prj\<^sub>0"
using trg_trg_f.prj_in_hom by auto
show "C.dom trg.leg0 = C.cod trg_trg_f.Prj\<^sub>0\<^sub>1"
by simp
show "trg.leg0 \<cdot> trg_trg_f.Prj\<^sub>0\<^sub>1 = f.leg1 \<cdot> trg_trg_f.Prj\<^sub>0"
by (metis C.comp_assoc C.prj0_simps_arr C.pullback_commutes'
\<open>C.span trg_trg_f.Prj\<^sub>0\<^sub>1 trg_trg_f.Prj\<^sub>0\<close>)
qed
moreover have "C.seq trg_trg_f.Prj\<^sub>0\<^sub>1 trg_trg_f.chine_assoc'"
by blast
ultimately show ?thesis
using C.comp_tuple_arr [of trg.leg0 f.leg1 trg_trg_f.Prj\<^sub>0\<^sub>1 trg_trg_f.Prj\<^sub>0
trg_trg_f.chine_assoc']
by auto
qed
also have "... = \<langle>trg_trg_f.Prj\<^sub>1\<^sub>0 \<lbrakk>trg.leg0, f.leg1\<rbrakk> trg_trg_f.Prj\<^sub>0\<^sub>0\<rangle>"
by simp
finally show ?thesis by simp
qed
finally show ?thesis by blast
qed
qed
qed
thus ?thesis by simp
qed
lemma runit'_ide_eq:
assumes "ide f"
shows "\<r>\<^sup>-\<^sup>1[f] = \<lparr>Chn = \<langle>Chn f \<lbrakk>Leg0 (Dom f), C.cod (Leg0 (Dom f))\<rbrakk> Leg0 (Dom f)\<rangle>,
Dom = Cod f,
Cod = \<lparr>Leg0 = \<p>\<^sub>0[Leg0 (Dom f), C.cod (Leg0 (Dom f))],
Leg1 = Leg1 (Dom f) \<cdot> \<p>\<^sub>1[Leg0 (Dom f), C.cod (Leg0 (Dom f))]\<rparr>\<rparr>"
proof -
interpret f: identity_arrow_of_spans C f
using assms ide_char' by auto
show "\<r>\<^sup>-\<^sup>1[f] = \<lparr>Chn = \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>, Dom = Cod f,
Cod = \<lparr>Leg0 = \<p>\<^sub>0[f.leg0, f.dsrc], Leg1 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.dsrc]\<rparr>\<rparr>"
proof -
have "C.inverse_arrows \<p>\<^sub>1[f.leg0, f.dsrc] \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using C.pullback_arr_cod(1) f.chine_eq_apex f.dom.apex_def f.dom.leg_simps(1)
by presburger
hence "C.inv \<p>\<^sub>1[f.leg0, f.dsrc] = \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using C.inv_is_inverse C.inverse_arrow_unique by auto
hence "\<r>\<^sup>-\<^sup>1[f] = \<lparr>Chn = \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>,
Dom = Cod \<r>[f], Cod = Dom \<r>[f]\<rparr>"
using assms runit_ide_eq inv_eq [of "\<r>[f]"] iso_runit by simp
thus ?thesis
using assms runit_ide_eq by simp
qed
qed
lemma lunit'_ide_eq:
assumes "ide f"
shows "\<l>\<^sup>-\<^sup>1[f] = \<lparr>Chn = \<langle>Leg1 (Dom f) \<lbrakk>C.cod (Leg1 (Dom f)), Leg1 (Dom f)\<rbrakk> Chn f\<rangle>,
Dom = Cod f,
Cod = \<lparr>Leg0 = Leg0 (Dom f) \<cdot> \<p>\<^sub>0[C.cod (Leg1 (Dom f)), Leg1 (Dom f)],
Leg1 = \<p>\<^sub>1[C.cod (Leg1 (Dom f)), Leg1 (Dom f)]\<rparr>\<rparr>"
proof -
interpret f: identity_arrow_of_spans C f
using assms ide_char' by auto
show "\<l>\<^sup>-\<^sup>1[f] = \<lparr>Chn = \<langle>f.leg1 \<lbrakk>f.dtrg, f.leg1\<rbrakk> f.chine\<rangle>, Dom = Cod f,
Cod = \<lparr>Leg0 = f.leg0 \<cdot> \<p>\<^sub>0[f.dtrg, f.leg1], Leg1 = \<p>\<^sub>1[f.dtrg, f.leg1]\<rparr>\<rparr>"
proof -
have "C.inverse_arrows \<p>\<^sub>0[f.dtrg, f.leg1] \<langle>f.leg1 \<lbrakk>f.dtrg, f.leg1\<rbrakk> f.chine\<rangle>"
using C.pullback_arr_cod(2) f.chine_eq_apex f.dom.apex_def f.dom.is_span
by presburger
hence "C.inv \<p>\<^sub>0[f.dtrg, f.leg1] = \<langle>f.leg1 \<lbrakk>f.dtrg, f.leg1\<rbrakk> f.chine\<rangle>"
using C.inv_is_inverse C.inverse_arrow_unique by auto
hence "\<l>\<^sup>-\<^sup>1[f] = \<lparr>Chn = \<langle>f.leg1 \<lbrakk>f.dtrg, f.leg1\<rbrakk> f.chine\<rangle>,
Dom = Cod \<l>[f], Cod = Dom \<l>[f]\<rparr>"
using assms lunit_ide_eq inv_eq [of "\<l>[f]"] iso_lunit by simp
thus ?thesis
using assms lunit_ide_eq by simp
qed
qed
end
locale adjunction_data_in_span_bicategory =
span_bicategory C prj0 prj1 +
adjunction_data_in_bicategory vcomp hcomp assoc unit src trg f g \<eta> \<epsilon>
for C :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<cdot>" 55)
and prj0 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<p>\<^sub>0[_, _]")
and prj1 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<p>\<^sub>1[_, _]")
and f :: "'a arrow_of_spans_data"
and g :: "'a arrow_of_spans_data"
and \<eta> :: "'a arrow_of_spans_data"
and \<epsilon> :: "'a arrow_of_spans_data"
begin
interpretation f: identity_arrow_of_spans C f
using ide_char' [of f] by auto
interpretation g: identity_arrow_of_spans C g
using ide_char' [of g] by auto
interpretation gf: two_composable_identity_arrows_of_spans C prj0 prj1 g f
using antipar by (unfold_locales, auto)
interpretation fg: two_composable_identity_arrows_of_spans C prj0 prj1 f g
using antipar by (unfold_locales, auto)
interpretation fgf: three_composable_identity_arrows_of_spans C prj0 prj1 f g f ..
interpretation gfg: three_composable_identity_arrows_of_spans C prj0 prj1 g f g ..
interpretation \<eta>: arrow_of_spans C \<eta>
using arr_char unit_in_hom by auto
interpretation \<epsilon>: arrow_of_spans C \<epsilon>
using arr_char counit_in_hom by auto
lemma chine_unit_in_hom:
shows "\<guillemotleft>\<eta>.chine : f.dsrc \<rightarrow>\<^sub>C g.leg0 \<down>\<down> f.leg1\<guillemotright>"
proof -
have "\<guillemotleft>\<eta>.chine : \<eta>.dom.apex \<rightarrow>\<^sub>C \<eta>.cod.apex\<guillemotright>"
using \<eta>.chine_in_hom by simp
moreover have "\<eta>.dom.apex = f.dsrc"
using \<eta>.dom.apex_def dom_char unit_simps src_def by auto
moreover have "\<eta>.cod.apex = g.leg0 \<down>\<down> f.leg1"
by (metis arrow_of_spans_data.select_convs(1) cod_char gf.chine_composite
unit_simps(1,3))
ultimately show ?thesis by simp
qed
lemma chine_counit_in_hom:
shows "\<guillemotleft>\<epsilon>.chine : f.leg0 \<down>\<down> g.leg1 \<rightarrow>\<^sub>C f.dtrg\<guillemotright>"
proof -
have "\<guillemotleft>\<epsilon>.chine : \<epsilon>.dom.apex \<rightarrow>\<^sub>C \<epsilon>.cod.apex\<guillemotright>"
using \<epsilon>.chine_in_hom by simp
moreover have "\<epsilon>.cod.apex = f.dtrg"
using \<epsilon>.cod.apex_def cod_char counit_simps trg_def gf.composable by auto
moreover have "\<epsilon>.dom.apex = f.leg0 \<down>\<down> g.leg1"
by (metis Chn_in_hom \<epsilon>.chine_simps(2) category.in_homE counit_in_hom(2)
fg.chine_composite span_vertical_category_axioms span_vertical_category_def)
ultimately show ?thesis by simp
qed
lemma \<eta>_leg_simps:
shows "\<eta>.dom.leg0 = f.dsrc" and "\<eta>.dom.leg1 = f.dsrc"
and "\<eta>.cod.leg0 = gf.leg0" and "\<eta>.cod.leg1 = gf.leg1"
proof -
show "\<eta>.dom.leg0 = f.dsrc"
using dom_char unit_simps(2) src_def by auto
show "\<eta>.dom.leg1 = f.dsrc"
using dom_char unit_simps(2) src_def by auto
show "\<eta>.cod.leg0 = gf.leg0"
using cod_char unit_simps(1,3)
by (metis (no_types, lifting) arrow_of_spans_data.select_convs(2))
show "\<eta>.cod.leg1 = gf.leg1"
using cod_char unit_simps(1,3)
by (metis (no_types, lifting) arrow_of_spans_data.select_convs(2))
qed
lemma \<epsilon>_leg_simps:
shows "\<epsilon>.cod.leg0 = f.dtrg" and "\<epsilon>.cod.leg1 = f.dtrg"
and "\<epsilon>.dom.leg0 = fg.leg0" and "\<epsilon>.dom.leg1 = fg.leg1"
proof -
show "\<epsilon>.cod.leg0 = f.dtrg"
using cod_char counit_simps(3) trg_def gf.composable by auto
show "\<epsilon>.cod.leg1 = f.dtrg"
using cod_char counit_simps(3) trg_def gf.composable by auto
show "\<epsilon>.dom.leg0 = fg.leg0"
using dom_char counit_simps hcomp_def fg.composable by simp
show "\<epsilon>.dom.leg1 = fg.leg1"
using dom_char counit_simps hcomp_def fg.composable by simp
qed
lemma Chn_triangle_eq:
shows "Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) = gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> f.leg0"
and "Chn (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) = gf.prj\<^sub>1 \<cdot> \<eta>.chine \<cdot> g.leg1"
proof -
have 1: "Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) =
\<p>\<^sub>0[f.dtrg, f.leg1] \<cdot> chine_hcomp \<epsilon> f \<cdot> fgf.chine_assoc' \<cdot> chine_hcomp f \<eta> \<cdot>
\<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
proof -
have "Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) =
Chn \<l>[f] \<cdot> Chn (\<epsilon> \<star> f) \<cdot> Chn \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> Chn (f \<star> \<eta>) \<cdot> Chn \<r>\<^sup>-\<^sup>1[f]"
using antipar Chn_vcomp by auto
also have "... = \<p>\<^sub>0[f.dtrg, f.leg1] \<cdot> chine_hcomp \<epsilon> f \<cdot> fgf.chine_assoc' \<cdot>
chine_hcomp f \<eta> \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using \<alpha>_ide fg.composable gf.composable fgf.chine_assoc_inverse
C.inverse_unique inv_eq iso_assoc lunit_ide_eq hcomp_def [of \<epsilon> f]
gf.composable hcomp_def [of f \<eta>] fg.composable runit'_ide_eq
by simp
finally show ?thesis by blast
qed
moreover have "C.arr (Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]))"
by (meson arrI arrow_of_spans.chine_simps(1) arr_char triangle_in_hom(3))
ultimately have 2: "C.arr (\<p>\<^sub>0[f.dtrg, f.leg1] \<cdot> chine_hcomp \<epsilon> f \<cdot> fgf.chine_assoc' \<cdot>
chine_hcomp f \<eta> \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>)"
by simp
have "Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) =
\<p>\<^sub>0[f.dtrg, f.leg1] \<cdot> chine_hcomp \<epsilon> f \<cdot> fgf.chine_assoc' \<cdot> chine_hcomp f \<eta> \<cdot>
\<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using 1 by simp
also have
3: "... = fgf.Prj\<^sub>0 \<cdot> fgf.chine_assoc' \<cdot> chine_hcomp f \<eta> \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
by (metis (no_types, lifting) C.comp_assoc C.comp_cod_arr \<epsilon>_leg_simps(1,3)
chine_hcomp_props(5) counit_simps(1,4) f.chine_eq_apex f.cod_simps(3)
fg.composite_is_arrow fg.leg0_composite fgf.prj_simps(3,9) hseqE)
also have 4: "... = fgf.Prj\<^sub>0\<^sub>0 \<cdot> chine_hcomp f \<eta> \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using C.comp_reduce [of fgf.Prj\<^sub>0 fgf.chine_assoc' fgf.Prj\<^sub>0\<^sub>0
"chine_hcomp f \<eta> \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"]
2 fgf.prj_chine_assoc'(3)
by blast
also have "... = (gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc]) \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
proof -
have "fgf.Prj\<^sub>0\<^sub>0 \<cdot> chine_hcomp f \<eta> = gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc]"
proof -
have "fgf.Prj\<^sub>0\<^sub>0 \<cdot> chine_hcomp f \<eta> =
(gf.prj\<^sub>0 \<cdot> \<p>\<^sub>0[f.leg0, gf.leg1]) \<cdot>
\<langle>\<p>\<^sub>1[f.leg0, f.dsrc] \<lbrakk>f.leg0, gf.leg1\<rbrakk> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc]\<rangle>"
using hcomp_def fg.composable gf.composable chine_hcomp_ide_arr \<eta>_leg_simps by auto
also have "... = gf.prj\<^sub>0 \<cdot> \<p>\<^sub>0[f.leg0, gf.leg1] \<cdot>
\<langle>\<p>\<^sub>1[f.leg0, f.dsrc] \<lbrakk>f.leg0, gf.leg1\<rbrakk> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc]\<rangle>"
using C.comp_assoc by simp
also have "... = gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc]"
proof -
have "C.commutative_square f.leg0 gf.leg1 \<p>\<^sub>1[f.leg0, f.dsrc] (\<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc])"
proof
show "C.cospan f.leg0 gf.leg1"
using hcomp_def gf.composable fgf.prj_in_hom(5) by auto
show "C.span \<p>\<^sub>1[f.leg0, f.dsrc] (\<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc])"
using chine_unit_in_hom by auto
show "C.dom f.leg0 = C.cod \<p>\<^sub>1[f.leg0, f.dsrc]" by simp
show "f.leg0 \<cdot> \<p>\<^sub>1[f.leg0, f.dsrc] = gf.leg1 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc]"
by (metis C.comp_assoc C.pullback_commutes' \<eta>.cod_trg_eq_dom_trg
\<eta>.dom.leg_simps(1) \<eta>.leg1_commutes \<eta>_leg_simps(1-2,4)
\<open>C.cospan f.leg0 gf.leg1\<close>)
qed
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
moreover have "C.seq fgf.Prj\<^sub>0\<^sub>0 (chine_hcomp f \<eta>)"
using chine_hcomp_props(1) by (metis 2 3 4 C.match_2 C.seqE)
moreover have "C.seq (chine_hcomp f \<eta>) \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using chine_hcomp_props(1) by (metis 2 C.seqE)
ultimately show ?thesis
using C.comp_reduce by simp
qed
also have "... = gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>0[f.leg0, f.dsrc] \<cdot> \<langle>f.chine \<lbrakk>f.leg0, f.dsrc\<rbrakk> f.leg0\<rangle>"
using C.comp_assoc by simp
also have "... = gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> f.leg0"
using C.comp_cod_arr f.leg0_commutes by simp
finally show "Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) = gf.prj\<^sub>0 \<cdot> \<eta>.chine \<cdot> f.leg0"
by simp
have 1: "Chn (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) =
\<p>\<^sub>1[g.leg0, g.dsrc] \<cdot> chine_hcomp g \<epsilon> \<cdot> gfg.chine_assoc \<cdot> chine_hcomp \<eta> g \<cdot>
\<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
proof -
have "Chn (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) =
Chn \<r>[g] \<cdot> Chn (g \<star> \<epsilon>) \<cdot> Chn \<a>[g, f, g] \<cdot> Chn (\<eta> \<star> g) \<cdot> Chn \<l>\<^sup>-\<^sup>1[g]"
using antipar Chn_vcomp by auto
also have "... = \<p>\<^sub>1[g.leg0, g.dsrc] \<cdot> chine_hcomp g \<epsilon> \<cdot> gfg.chine_assoc \<cdot> chine_hcomp \<eta> g \<cdot>
\<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
using \<alpha>_ide gf.composable fg.composable runit_ide_eq hcomp_def [of g \<epsilon>]
fg.composable hcomp_def [of \<eta> g] gf.composable lunit'_ide_eq
by simp
finally show ?thesis by blast
qed
hence 2: "C.arr (\<p>\<^sub>1[g.leg0, g.dsrc] \<cdot> chine_hcomp g \<epsilon> \<cdot> gfg.chine_assoc \<cdot>
chine_hcomp \<eta> g \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>)"
by (metis arrI arr_char arrow_of_spans.chine_simps(1) triangle_in_hom(4))
have "Chn (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) =
\<p>\<^sub>1[g.leg0, g.dsrc] \<cdot> chine_hcomp g \<epsilon> \<cdot> gfg.chine_assoc \<cdot> chine_hcomp \<eta> g \<cdot>
\<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
using 1 by simp
also have
"... = gfg.Prj\<^sub>1 \<cdot> gfg.chine_assoc \<cdot> chine_hcomp \<eta> g \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
proof -
have "\<p>\<^sub>1[g.leg0, g.dsrc] \<cdot> chine_hcomp g \<epsilon> = \<p>\<^sub>1[g.leg0, \<epsilon>.dom.leg1]"
by (metis C.comp_cod_arr \<epsilon>_leg_simps(2) \<epsilon>_leg_simps(4) arrI chine_hcomp_props(6)
fg.leg1_composite g.cod_simps(2) g.identity_arrow_of_spans_axioms
gfg.cospan_\<mu>\<nu> gfg.prj_simps(10) gfg.prj_simps(16) hseqE
identity_arrow_of_spans.chine_eq_apex seqE triangle_in_hom(4))
also have "... = gfg.Prj\<^sub>1"
using dom_char counit_simps hcomp_def fg.composable by simp
finally have "\<p>\<^sub>1[g.leg0, g.dsrc] \<cdot> chine_hcomp g \<epsilon> = gfg.Prj\<^sub>1"
by simp
moreover have "C.seq \<p>\<^sub>1[g.leg0, g.dsrc] (chine_hcomp g \<epsilon>)"
using chine_hcomp_props(1) [of g \<epsilon>] gf.composable calculation gfg.prj_in_hom(4)
by auto
moreover have "C.seq (chine_hcomp g \<epsilon>)
(gfg.chine_assoc \<cdot> chine_hcomp \<eta> g \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>)"
using chine_hcomp_props(1) gf.composable 2 by (metis C.seqE)
ultimately show ?thesis
using C.comp_reduce by simp
qed
also have "... = gfg.Prj\<^sub>1\<^sub>1 \<cdot> chine_hcomp \<eta> g \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
using C.comp_reduce [of gfg.Prj\<^sub>1 gfg.chine_assoc gfg.Prj\<^sub>1\<^sub>1
"chine_hcomp \<eta> g \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"]
2 gfg.prj_chine_assoc(1)
by blast
also have "... = (gf.prj\<^sub>1 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>1[g.dtrg, g.leg1]) \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
by (metis C.comp_assoc \<eta>_leg_simps(1) \<eta>_leg_simps(3) arrI chine_hcomp_props(6)
g.cod_simps(3) gf.leg0_composite gfg.cospan_\<nu>\<pi> hseqE seqE triangle_in_hom(4))
also have "... = gf.prj\<^sub>1 \<cdot> \<eta>.chine \<cdot> \<p>\<^sub>1[g.dtrg, g.leg1] \<cdot> \<langle>g.leg1 \<lbrakk>g.dtrg, g.leg1\<rbrakk> g.chine\<rangle>"
using C.comp_assoc by simp
also have "... = gf.prj\<^sub>1 \<cdot> \<eta>.chine \<cdot> g.leg1"
using C.comp_cod_arr g.leg1_commutes by simp
finally show "Chn (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) = gf.prj\<^sub>1 \<cdot> \<eta>.chine \<cdot> g.leg1"
by simp
qed
end
subsection "Maps in Span(C)"
text \<open>
In this section, we chararacterize the maps (\emph{i.e}~the left adjoints)
in a span bicategory. This is Proposition 2 of \cite{carboni-et-al}.
\<close>
context span_bicategory
begin
abbreviation adjoint_of_map
where "adjoint_of_map f \<equiv> \<lparr>Chn = Chn f,
Dom = \<lparr>Leg0 = Leg1 (Dom f), Leg1 = Leg0 (Dom f)\<rparr>,
Cod = \<lparr>Leg0 = Leg1 (Dom f), Leg1 = Leg0 (Dom f)\<rparr>\<rparr>"
abbreviation unit_for_map
where "unit_for_map f \<equiv> \<lparr>Chn = \<langle>C.inv (Leg0 (Dom f))
\<lbrakk>Leg1 (Dom f), Leg1 (Dom f)\<rbrakk>
C.inv (Leg0 (Dom f))\<rangle>,
Dom = Dom (src f),
Cod = Dom (hcomp (adjoint_of_map f) f)\<rparr>"
abbreviation counit_for_map
where "counit_for_map f \<equiv> \<lparr>Chn = Leg1 (Dom f) \<cdot> \<p>\<^sub>0[Leg0 (Dom f), Leg0 (Dom f)],
Dom = Dom (hcomp f (adjoint_of_map f)),
Cod = Dom (trg f)\<rparr>"
lemma is_left_adjoint_char:
shows "is_left_adjoint f \<longleftrightarrow> ide f \<and> C.iso (Leg0 (Dom f))"
and "is_left_adjoint f \<Longrightarrow>
adjunction_in_bicategory vcomp hcomp assoc unit src trg f
(adjoint_of_map f) (unit_for_map f) (counit_for_map f)"
proof
show 1: "is_left_adjoint f \<Longrightarrow> ide f \<and> C.iso (Leg0 (Dom f))"
proof
assume f: "is_left_adjoint f"
obtain g \<eta> \<epsilon> where adj: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f g \<eta> \<epsilon>"
using f adjoint_pair_def by blast
interpret adjunction_in_bicategory vcomp hcomp assoc unit src trg f g \<eta> \<epsilon>
using adj by auto
show "ide f" by simp
interpret f: identity_arrow_of_spans C f
using ide_char' [of f] by auto
interpret g: identity_arrow_of_spans C g
using ide_char' [of g] by auto
interpret gf: two_composable_identity_arrows_of_spans C prj0 prj1 g f
using antipar by (unfold_locales, auto)
interpret fg: two_composable_identity_arrows_of_spans C prj0 prj1 f g
using antipar by (unfold_locales, auto)
interpret fgf: three_composable_identity_arrows_of_spans C prj0 prj1 f g f ..
interpret src_f: arrow_of_spans C \<open>src f\<close>
using arr_char gf.are_arrows(2) by blast
interpret src_f: identity_arrow_of_spans C \<open>src f\<close>
using ide_char ide_src src_def by (unfold_locales, simp)
interpret \<eta>: arrow_of_spans C \<eta>
using arr_char unit_in_hom by auto
interpret \<epsilon>: arrow_of_spans C \<epsilon>
using arr_char counit_in_hom by auto
interpret adjunction_data_in_span_bicategory C prj0 prj1 f g \<eta> \<epsilon>
..
show "C.iso f.leg0"
proof -
have "C.section f.leg0"
proof -
have "f.chine = (gf.prj\<^sub>0 \<cdot> \<eta>.chine) \<cdot> f.leg0"
using triangle_left' Chn_triangle_eq(1) C.comp_assoc by simp
thus ?thesis
using f.chine_is_identity by auto
qed
moreover have "C.retraction f.leg0"
using C.retractionI [of f.leg0 "gf.prj\<^sub>0 \<cdot> \<eta>.chine"] hcomp_def C.comp_assoc
\<eta>.leg0_commutes gf.leg0_composite \<eta>_leg_simps
by auto
ultimately show ?thesis
by (simp add: C.iso_iff_section_and_retraction)
qed
qed
have 2: "ide f \<and> C.iso (Leg0 (Dom f)) \<Longrightarrow>
adjunction_in_bicategory vcomp hcomp assoc unit src trg f
(adjoint_of_map f) (unit_for_map f) (counit_for_map f)"
text \<open>
The right adjoint \<open>g\<close> is obtained by exchanging the legs of \<open>f\<close>.
The unit is obtained by tupling \<open>C.inv f.leg0\<close> with itself,
via the pullback of \<open>f.leg1\<close> with itself.
The counit is given by the legs of \<open>f \<star> g\<close>, which are equal,
because the two legs of a pullback of the isomorphism \<open>f.leg0\<close>
with itself must be equal.
It then remains to verify the triangle identities.
\<close>
proof -
assume f: "ide f \<and> C.iso (Leg0 (Dom f))"
interpret f: identity_arrow_of_spans C f
using f ide_char' by auto
interpret Dom_src: span_in_category C \<open>\<lparr>Leg0 = f.dsrc, Leg1 = f.dsrc\<rparr>\<close>
using f by (unfold_locales, auto)
interpret Dom_trg: span_in_category C \<open>\<lparr>Leg0 = f.dtrg, Leg1 = f.dtrg\<rparr>\<close>
using f by (unfold_locales, auto)
define g where "g = adjoint_of_map f"
interpret Dom_g: span_in_category C \<open>\<lparr>Leg0 = f.leg1, Leg1 = f.leg0\<rparr>\<close>
by (unfold_locales, simp)
interpret g: arrow_of_spans C g
unfolding g_def
using Dom_g.apex_def f.leg0_commutes f.leg1_commutes
by (unfold_locales, auto)
interpret g: identity_arrow_of_spans C g
using g_def
by (unfold_locales, auto)
have ide_g: "ide g"
using f ide_char g.arrow_of_spans_axioms by simp
interpret fg: two_composable_arrows_of_spans C prj0 prj1 f g
apply unfold_locales
using g_def src_def trg_def arr_char f.arrow_of_spans_axioms g.arrow_of_spans_axioms
by auto
interpret fg: two_composable_identity_arrows_of_spans C prj0 prj1 f g
..
interpret gf: two_composable_arrows_of_spans C prj0 prj1 g f
apply unfold_locales
using g_def src_def trg_def arr_char f.arrow_of_spans_axioms g.arrow_of_spans_axioms
by auto
interpret gf: two_composable_identity_arrows_of_spans C prj0 prj1 g f
..
have hcomp_fg_eq: "hcomp f g = \<lparr>Chn = f.leg0 \<down>\<down> f.leg0,
Dom = \<lparr>Leg0 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.leg0],
Leg1 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.leg0]\<rparr>,
Cod = \<lparr>Leg0 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.leg0],
Leg1 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.leg0]\<rparr>\<rparr>"
using f g_def hcomp_def fg.composable src_def trg_def arr_char f.arrow_of_spans_axioms
g.arrow_of_spans_axioms chine_hcomp_def gf.are_identities(1) chine_hcomp_ide_ide
C.pullback_iso_self
by auto
have hcomp_gf_eq: "hcomp g f = \<lparr>Chn = f.leg1 \<down>\<down> f.leg1,
Dom = \<lparr>Leg0 = f.leg0 \<cdot> \<p>\<^sub>0[f.leg1, f.leg1],
Leg1 = f.leg0 \<cdot> \<p>\<^sub>1[f.leg1, f.leg1]\<rparr>,
Cod = \<lparr>Leg0 = f.leg0 \<cdot> \<p>\<^sub>0[f.leg1, f.leg1],
Leg1 = f.leg0 \<cdot> \<p>\<^sub>1[f.leg1, f.leg1]\<rparr>\<rparr>"
using g_def hcomp_def gf.composable src_def trg_def chine_hcomp_ide_ide
arr_char f.arrow_of_spans_axioms g.arrow_of_spans_axioms ide_char
by simp
define \<eta> where "\<eta> = unit_for_map f"
interpret Dom_gf: span_in_category C \<open>\<lparr>Leg0 = f.leg0 \<cdot> \<p>\<^sub>0[f.leg1, f.leg1],
Leg1 = f.leg0 \<cdot> \<p>\<^sub>1[f.leg1, f.leg1]\<rparr>\<close>
by (unfold_locales, simp_all)
interpret \<eta>: arrow_of_spans C \<eta>
using f g_def \<eta>_def hcomp_def src_def trg_def f.arrow_of_spans_axioms
g.arrow_of_spans_axioms arr_char C.comp_arr_inv'
C.tuple_in_hom [of f.leg1 f.leg1 "C.inv f.leg0" "C.inv f.leg0"]
Dom_src.apex_def Dom_gf.apex_def
apply unfold_locales by (simp_all add: C.comp_assoc)
have unit_in_hom: "\<guillemotleft>\<eta> : src f \<Rightarrow> hcomp g f\<guillemotright>"
proof
show 1: "arr \<eta>"
using arr_char \<eta>.arrow_of_spans_axioms by simp
show "dom \<eta> = src f"
using 1 \<eta>_def dom_char src_def Dom_src.apex_def by simp
show "cod \<eta> = hcomp g f"
using 1 \<eta>_def g_def cod_char hcomp_gf_eq Dom_gf.apex_def by simp
qed
define \<epsilon> where "\<epsilon> = counit_for_map f"
interpret Dom_fg: span_in_category C \<open>\<lparr>Leg0 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.leg0],
Leg1 = f.leg1 \<cdot> \<p>\<^sub>1[f.leg0, f.leg0]\<rparr>\<close>
by (unfold_locales, simp_all)
interpret \<epsilon>: arrow_of_spans C \<epsilon>
using f g_def \<epsilon>_def hcomp_def src_def trg_def f.arrow_of_spans_axioms
g.arrow_of_spans_axioms arr_char C.comp_cod_arr C.pullback_iso_self
Dom_trg.apex_def Dom_fg.apex_def
apply unfold_locales by auto
have counit_in_hom: "\<guillemotleft>\<epsilon> : hcomp f g \<Rightarrow> trg f\<guillemotright>"
proof
show 1: "arr \<epsilon>"
using arr_char \<epsilon>.arrow_of_spans_axioms by simp
show "cod \<epsilon> = trg f"
using 1 \<epsilon>_def cod_char trg_def Dom_trg.apex_def by simp
show "dom \<epsilon> = hcomp f g"
using 1 g_def \<epsilon>_def dom_char hcomp_fg_eq Dom_fg.apex_def by simp
qed
interpret adj: adjunction_data_in_bicategory vcomp hcomp assoc unit src trg f g \<eta> \<epsilon>
using f ide_g unit_in_hom counit_in_hom gf.composable
by (unfold_locales, simp_all)
interpret adjunction_data_in_span_bicategory C prj0 prj1 f g \<eta> \<epsilon>
..
have triangle_left: "(\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<bullet> \<r>[f]"
proof -
have "\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f] = f"
proof (intro arr_eqI)
show "par (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) f"
using f ide_in_hom [of f] adj.triangle_in_hom(3)
by (metis (no_types, lifting) in_homE)
show "Chn (\<l>[f] \<bullet> (\<epsilon> \<star> f) \<bullet> \<a>\<^sup>-\<^sup>1[f, g, f] \<bullet> (f \<star> \<eta>) \<bullet> \<r>\<^sup>-\<^sup>1[f]) = f.chine"
using f g_def \<eta>_def Chn_triangle_eq(1) C.comp_tuple_arr C.comp_inv_arr' by simp
qed
thus ?thesis
using adj.triangle_equiv_form by simp
qed
have triangle_right: "(g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<bullet> \<l>[g]"
proof -
have "\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g] = g"
proof (intro arr_eqI)
show "par (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) g"
using adj.ide_right ide_in_hom [of g] adj.triangle_in_hom(4)
by (metis (no_types, lifting) in_homE)
show "Chn (\<r>[g] \<bullet> (g \<star> \<epsilon>) \<bullet> \<a>[g, f, g] \<bullet> (\<eta> \<star> g) \<bullet> \<l>\<^sup>-\<^sup>1[g]) = g.chine"
using f g_def \<eta>_def Chn_triangle_eq(2) C.comp_tuple_arr C.comp_inv_arr' by simp
qed
thus ?thesis
using adj.triangle_equiv_form by simp
qed
interpret adj: adjunction_in_bicategory vcomp hcomp assoc unit src trg f g \<eta> \<epsilon>
using triangle_left triangle_right by (unfold_locales, simp_all)
show "adjunction_in_bicategory vcomp hcomp assoc unit src trg f g \<eta> \<epsilon>" ..
qed
show "ide f \<and> C.iso (Leg0 (Dom f)) \<Longrightarrow> is_left_adjoint f"
using 2 adjoint_pair_def by blast
show "is_left_adjoint f \<Longrightarrow> adjunction_in_bicategory vcomp hcomp assoc unit src trg f
(adjoint_of_map f) (unit_for_map f) (counit_for_map f)"
using 1 2 by blast
qed
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,4210 +1,4210 @@
(* Title: Strictness
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Strictness"
theory Strictness
imports Category3.ConcreteCategory Pseudofunctor CanonicalIsos
begin
text \<open>
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 \<open>B\<close> we may construct a related strict bicategory \<open>S\<close>,
its \emph{strictification}, together with a pseudofunctor that embeds \<open>B\<close> in \<open>S\<close>.
The Strictness Theorem states that this pseudofunctor is an equivalence pseudofunctor,
so that bicategory \<open>B\<close> 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.
\<close>
subsection "Normal and Strict Bicategories"
text \<open>
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''.
\<close>
locale normal_bicategory =
bicategory +
assumes strict_lunit: "\<And>f. ide f \<Longrightarrow> \<l>[f] = f"
and strict_runit: "\<And>f. ide f \<Longrightarrow> \<r>[f] = f"
begin
lemma strict_unit:
assumes "obj a"
shows "ide \<i>[a]"
using assms strict_runit unitor_coincidence(2) [of a] by auto
lemma strict_lunit':
assumes "ide f"
shows "\<l>\<^sup>-\<^sup>1[f] = f"
using assms strict_lunit by simp
lemma strict_runit':
assumes "ide f"
shows "\<r>\<^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 \<star> 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 \<star> a = f"
using assms strict_runit
by (metis comp_arr_dom comp_ide_arr ide_cod ide_dom runit_naturality)
end
text \<open>
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''.
\<close>
locale strict_bicategory =
normal_bicategory +
assumes strict_assoc: "\<And>f g h. \<lbrakk>ide f; ide g; ide h; src f = trg g; src g = trg h\<rbrakk> \<Longrightarrow>
ide \<a>[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 \<a>\<^sup>-\<^sup>1[f, g, h]"
using assms strict_assoc by simp
lemma hcomp_assoc:
shows "(\<mu> \<star> \<nu>) \<star> \<tau> = \<mu> \<star> \<nu> \<star> \<tau>"
proof (cases "hseq \<mu> \<nu> \<and> hseq \<nu> \<tau>")
show "\<not> (hseq \<mu> \<nu> \<and> hseq \<nu> \<tau>) \<Longrightarrow> ?thesis"
by (metis hseqE hseq_char' match_1 match_2)
show "hseq \<mu> \<nu> \<and> hseq \<nu> \<tau> \<Longrightarrow> ?thesis"
proof -
assume 1: "hseq \<mu> \<nu> \<and> hseq \<nu> \<tau>"
have 2: "arr \<mu> \<and> arr \<nu> \<and> arr \<tau> \<and> src \<mu> = trg \<nu> \<and> src \<nu> = trg \<tau>"
using 1 by blast
have "(\<mu> \<star> \<nu>) \<star> \<tau> = \<a>[cod \<mu>, cod \<nu>, cod \<tau>] \<cdot> ((\<mu> \<star> \<nu>) \<star> \<tau>)"
using 1 assoc_in_hom strict_assoc comp_cod_arr assoc_simps(4) hseq_char
by simp
also have "... = (\<mu> \<star> \<nu> \<star> \<tau>) \<cdot> \<a>[dom \<mu>, dom \<nu>, dom \<tau>]"
using 1 assoc_naturality by auto
also have "... = \<mu> \<star> \<nu> \<star> \<tau>"
using 2 assoc_in_hom [of "dom \<mu>" "dom \<nu>" "dom \<tau>"] strict_assoc comp_arr_dom
by auto
finally show ?thesis by simp
qed
qed
text \<open>
In a strict bicategory, every canonical isomorphism is an identity.
\<close>
interpretation bicategorical_language ..
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
lemma ide_eval_Can:
assumes "Can t"
shows "ide \<lbrace>t\<rbrace>"
proof -
have 1: "\<And>u1 u2. \<lbrakk> ide \<lbrace>u1\<rbrace>; ide \<lbrace>u2\<rbrace>; Arr u1; Arr u2; Dom u1 = Cod u2 \<rbrakk>
\<Longrightarrow> ide (\<lbrace>u1\<rbrace> \<cdot> \<lbrace>u2\<rbrace>)"
by (metis (no_types, lifting) E.eval_simps'(4-5) comp_ide_self ide_char)
have "\<And>u. Can u \<Longrightarrow> ide \<lbrace>u\<rbrace>"
proof -
fix u
show "Can u \<Longrightarrow> ide \<lbrace>u\<rbrace>"
(* TODO: Rename \<ll>_ide_simp \<rr>_ide_simp to \<ll>_ide_eq \<rr>_ide_eq *)
using 1 \<alpha>_def \<a>'_def strict_lunit strict_runit strict_assoc strict_assoc'
\<ll>_ide_simp \<rr>_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>\<lfloor>f\<^bold>\<rfloor> = \<^bold>\<lfloor>g\<^bold>\<rfloor>"
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 \<open>
The following result gives conditions for strictness of a bicategory that are typically
somewhat easier to verify than those used for the definition.
\<close>
lemma is_strict_if:
assumes "\<And>f. ide f \<Longrightarrow> f \<star> src f = f"
and "\<And>f. ide f \<Longrightarrow> trg f \<star> f = f"
and "\<And>a. obj a \<Longrightarrow> ide \<i>[a]"
and "\<And>f g h. \<lbrakk>ide f; ide g; ide h; src f = trg g; src g = trg h\<rbrakk> \<Longrightarrow> ide \<a>[f, g, h]"
shows "strict_bicategory V H \<a> \<i> src trg"
proof
show "\<And>f g h. \<lbrakk>ide f; ide g; ide h; src f = trg g; src g = trg h\<rbrakk> \<Longrightarrow> ide \<a>[f, g, h]"
by fact
fix f
assume f: "ide f"
show "\<l>[f] = f"
proof -
have "f = \<l>[f]"
using assms f unit_simps(5)
by (intro lunit_eqI) (auto simp add: comp_arr_ide)
thus ?thesis by simp
qed
show "\<r>[f] = f"
proof -
have "f = \<r>[f]"
proof (intro runit_eqI)
show "ide f" by fact
show "\<guillemotleft>f : f \<star> src f \<Rightarrow> f\<guillemotright>"
using f assms(1) by auto
show "f \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f]"
proof -
have "(f \<star> \<i>[src f]) \<cdot> \<a>[f, src f, src f] = (f \<star> src f) \<cdot> \<a>[f, src f, src f]"
using f assms(2-3) unit_simps(5) by simp
also have "... = (f \<star> src f \<star> src f) \<cdot> \<a>[f, src f, src f]"
using f assms(1-2) ideD(1) trg_src src.preserves_ide by metis
also have "... = f \<star> 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 \<open>
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 \<open>P\<close> of bicategories that is ``bicategorical''
(\emph{i.e.}~respects biequivalence), if we want to show that \<open>P\<close> holds for a bicategory \<open>B\<close>
then it suffices to show that \<open>P\<close> holds for the strictification of \<open>B\<close>, and if we want to show
that \<open>P\<close> 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 \<open>B\<close>, the strictification \<open>S\<close> of \<open>B\<close> may be constructed as the bicategory
whose arrows are triples \<open>(A, B, \<mu>)\<close>, where \<open>X\<close> and \<open>Y\<close> are ``normal identity terms''
(essentially, nonempty horizontally composable lists of 1-cells of \<open>B\<close>) having the same
syntactic source and target, and \<open>\<guillemotleft>\<mu> : \<lbrace>X\<rbrace> \<Rightarrow> \<lbrace>Y\<rbrace>\<guillemotright>\<close> in \<open>B\<close>.
Vertical composition in \<open>S\<close> is given by composition of the underlying arrows in \<open>B\<close>.
Horizontal composition in \<open>S\<close> is given by \<open>(A, B, \<mu>) \<star> (A', B', \<mu>') = (AA', BB', \<nu>)\<close>,
where \<open>AA'\<close> and \<open>BB'\<close> denote concatenations of lists and where \<open>\<nu>\<close> is defined as the
composition \<open>can BB' (B \<^bold>\<star> B') \<cdot> (\<mu> \<star> \<mu>') \<cdot> can (A \<^bold>\<star> A') AA'\<close>, where \<open>can (A \<^bold>\<star> A') AA'\<close> and
\<open>can BB' (B \<^bold>\<star> B')\<close> are canonical isomorphisms in \<open>B\<close>. The canonical isomorphism
\<open>can (A \<^bold>\<star> A') AA'\<close> corresponds to taking a pair of lists \<open>A \<^bold>\<star> A'\<close> and
``shifting the parentheses to the right'' to obtain a single list \<open>AA'\<close>.
The canonical isomorphism can \<open>BB' (B \<^bold>\<star> B')\<close> corresponds to the inverse rearrangement.
The bicategory \<open>B\<close> embeds into its strictification \<open>S\<close> via the functor \<open>UP\<close> that takes
each arrow \<open>\<mu>\<close> of \<open>B\<close> to \<open>(\<^bold>\<langle>dom \<mu>\<^bold>\<rangle>, \<^bold>\<langle>cod \<mu>\<^bold>\<rangle>, \<mu>)\<close>, where \<open>\<^bold>\<langle>dom \<mu>\<^bold>\<rangle>\<close> and \<open>\<^bold>\<langle>cod \<mu>\<^bold>\<rangle>\<close> denote
one-element lists. This mapping extends to a pseudofunctor.
There is also a pseudofunctor \<open>DN\<close>, which maps \<open>(A, B, \<mu>)\<close> in \<open>S\<close> to \<open>\<mu>\<close> in \<open>B\<close>;
this is such that \<open>DN o UP\<close> is the identity on \<open>B\<close> and \<open>UP o DN\<close> is equivalent to the
identity on \<open>S\<close>, so we obtain a biequivalence between \<open>B\<close> and \<open>S\<close>.
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.
\<close>
locale strictified_bicategory =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
for V\<^sub>B :: "'a comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'a \<Rightarrow> 'a" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
begin
sublocale E: self_evaluation_map V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B ..
notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>B _\<guillemotright>")
notation E.eval ("\<lbrace>_\<rbrace>")
notation E.Nmlize ("\<^bold>\<lfloor>_\<^bold>\<rfloor>")
text \<open>
The following gives the construction of a bicategory whose arrows are triples \<open>(A, B, \<mu>)\<close>,
where \<open>Nml A \<and> Ide A\<close>, \<open>Nml B \<and> Ide B\<close>, \<open>Src A = Src B\<close>, \<open>Trg A = Trg B\<close>, and \<open>\<mu> : \<lbrace>A\<rbrace> \<Rightarrow> \<lbrace>B\<rbrace>\<close>.
We use @{locale concrete_category} to construct the vertical composition, so formally the
arrows of the bicategory will be of the form \<open>MkArr A B \<mu>\<close>.
\<close>
text \<open>
The 1-cells of the bicategory correspond to normal, identity terms \<open>A\<close>
in the bicategorical language associated with \<open>B\<close>.
\<close>
abbreviation IDE
where "IDE \<equiv> {A. E.Nml A \<and> E.Ide A}"
text \<open>
If terms \<open>A\<close> and \<open>B\<close> determine 1-cells of the strictification and have a
common source and target, then the 2-cells between these 1-cells correspond
to arrows \<open>\<mu>\<close> of the underlying bicategory such that \<open>\<guillemotleft>\<mu> : \<lbrace>A\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>B\<rbrace>\<guillemotright>\<close>.
\<close>
abbreviation HOM
where "HOM A B \<equiv> {\<mu>. E.Src A = E.Src B \<and> E.Trg A = E.Trg B \<and> \<guillemotleft>\<mu> : \<lbrace>A\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>B\<rbrace>\<guillemotright>}"
text \<open>
The map taking term \<open>A \<in> OBJ\<close> to its evaluation \<open>\<lbrace>A\<rbrace> \<in> HOM A A\<close> defines the
embedding of 1-cells as identity 2-cells.
\<close>
abbreviation EVAL
where "EVAL \<equiv> E.eval"
sublocale concrete_category IDE HOM EVAL \<open>\<lambda>_ _ _ \<mu> \<nu>. \<mu> \<cdot>\<^sub>B \<nu>\<close>
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 (\<lambda>_ _ _ \<mu> \<nu>. \<mu> \<cdot>\<^sub>B \<nu>)"
..
abbreviation vcomp (infixr "\<cdot>" 55)
where "vcomp \<equiv> COMP"
lemma arr_char:
shows "arr F \<longleftrightarrow>
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>
\<guillemotleft>Map F : \<lbrace>Dom F\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod F\<rbrace>\<guillemotright> \<and> F \<noteq> 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 "\<guillemotleft>Map F : \<lbrace>Dom F\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod F\<rbrace>\<guillemotright>" and "F \<noteq> Null"
shows "arr F"
using assms arr_char by blast
lemma arrE [elim]:
assumes "arr F"
shows "(\<lbrakk> 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);
\<guillemotleft>Map F : \<lbrace>Dom F\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod F\<rbrace>\<guillemotright>; F \<noteq> Null \<rbrakk> \<Longrightarrow> T) \<Longrightarrow> T"
using assms arr_char by simp
lemma ide_char:
shows "ide F \<longleftrightarrow> endo F \<and> B.ide (Map F)"
proof
show "ide F \<Longrightarrow> endo F \<and> B.ide (Map F)"
using ide_char by (simp add: E.ide_eval_Ide)
show "endo F \<and> B.ide (Map F) \<Longrightarrow> 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 "(\<lbrakk> arr F; Dom F = Cod F; B.ide (Map F); Map F = \<lbrace>Dom F\<rbrace>;
Map F = \<lbrace>Cod F\<rbrace> \<rbrakk> \<Longrightarrow> T) \<Longrightarrow> T"
using assms
by (metis (no_types, lifting) Map_ide(2) ide_char seq_char)
text \<open>
Source and target are defined by the corresponding syntactic operations on terms.
\<close>
definition src
where "src F \<equiv> if arr F then MkIde (E.Src (Dom F)) else null"
definition trg
where "trg F \<equiv> 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) = \<lbrace>E.Src (Dom F)\<rbrace>"
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) = \<lbrace>E.Trg (Dom F)\<rbrace>"
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 "\<And>g f. seq g f \<Longrightarrow> src (g \<cdot> f) = src g \<cdot> src f"
proof -
fix g f
assume gf: "seq g f"
have "src (g \<cdot> f) = MkIde (E.Src (Dom (g \<cdot> 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)) \<cdot> MkIde (E.Src (Dom f))"
using gf seq_char E.Obj_implies_Ide by auto
also have "... = src g \<cdot> src f"
using gf src_def comp_char by auto
finally show "src (g \<cdot> f) = src g \<cdot> 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 "\<And>g f. seq g f \<Longrightarrow> trg (g \<cdot> f) = trg g \<cdot> trg f"
proof -
fix g f
assume gf: "seq g f"
have "trg (g \<cdot> f) = MkIde (E.Trg (Dom (g \<cdot> 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)) \<cdot> MkIde (E.Trg (Dom f))"
using gf seq_char E.Obj_implies_Ide by auto
also have "... = trg g \<cdot> trg f"
using gf trg_def comp_char by auto
finally show "trg (g \<cdot> f) = trg g \<cdot> 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 ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
definition hcomp (infixr "\<star>" 53)
where "\<mu> \<star> \<nu> \<equiv> if arr \<mu> \<and> arr \<nu> \<and> src \<mu> = trg \<nu>
then MkArr (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>)
(B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>))
else null"
lemma arr_hcomp:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "arr (\<mu> \<star> \<nu>)"
proof -
have 1: "E.Ide (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<and> E.Nml (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<and>
E.Ide (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<and> E.Nml (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>)"
using assms arr_char src_def trg_def E.Ide_HcompNml E.Nml_HcompNml(1) by auto
moreover
have "\<guillemotleft>B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<cdot>\<^sub>B (Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) :
\<lbrace>Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>\<rbrace>\<guillemotright>"
proof -
have "\<guillemotleft>Map \<mu> \<star>\<^sub>B Map \<nu> : \<lbrace>Dom \<mu> \<^bold>\<star> Dom \<nu>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod \<mu> \<^bold>\<star> Cod \<nu>\<rbrace>\<guillemotright>"
using assms arr_char dom_char cod_char src_def trg_def E.eval_simps'(2-3)
by auto
moreover
have "\<guillemotleft>B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) :
\<lbrace>Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Dom \<mu> \<^bold>\<star> Dom \<nu>\<rbrace>\<guillemotright> \<and>
\<guillemotleft>B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) :
\<lbrace>Cod \<mu> \<^bold>\<star> Cod \<nu>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>\<rbrace>\<guillemotright>"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) = E.Src (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<and>
E.Trg (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) = E.Trg (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>)"
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 \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "src (\<mu> \<star> \<nu>) = src \<nu>"
using assms arr_char hcomp_def src_def trg_def arr_hcomp E.Src_HcompNml by simp
lemma trg_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "trg (hcomp \<mu> \<nu>) = trg \<mu>"
using assms arr_char hcomp_def src_def trg_def arr_hcomp E.Trg_HcompNml by simp
lemma hseq_char:
shows "arr (\<mu> \<star> \<nu>) \<longleftrightarrow> arr \<mu> \<and> arr \<nu> \<and> src \<mu> = trg \<nu>"
using arr_hcomp hcomp_def by simp
lemma Dom_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "Dom (\<mu> \<star> \<nu>) = Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>"
using assms hcomp_def [of \<mu> \<nu>] by simp
lemma Cod_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "Cod (\<mu> \<star> \<nu>) = Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>"
using assms hcomp_def [of \<mu> \<nu>] by simp
lemma Map_hcomp [simp]:
assumes "arr \<mu>" and "arr \<nu>" and "src \<mu> = trg \<nu>"
shows "Map (\<mu> \<star> \<nu>) = B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>)"
using assms hcomp_def [of \<mu> \<nu>] by simp
interpretation "functor" VV.comp vcomp \<open>\<lambda>\<mu>\<nu>. hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
proof
show "\<And>f. \<not> VV.arr f \<Longrightarrow> fst f \<star> snd f = null"
using hcomp_def by auto
show A: "\<And>f. VV.arr f \<Longrightarrow> arr (fst f \<star> snd f)"
using VV.arrE hseq_char by blast
show "\<And>f. VV.arr f \<Longrightarrow> dom (fst f \<star> snd f) = fst (VV.dom f) \<star> snd (VV.dom f)"
proof -
fix f
assume f: "VV.arr f"
have "dom (fst f \<star> snd f) = MkIde (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))"
using f VV.arrE [of f] dom_char arr_hcomp hcomp_def by simp
also have "... = fst (VV.dom f) \<star> snd (VV.dom f)"
proof -
have "hcomp (fst (VV.dom f)) (snd (VV.dom f)) =
MkArr (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))
(B.can (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)) (Dom (fst f) \<^bold>\<star> Dom (snd f)) \<cdot>\<^sub>B
(\<lbrace>Dom (fst f)\<rbrace> \<star>\<^sub>B \<lbrace>Dom (snd f)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)) (Dom (fst f) \<^bold>\<star> Dom (snd f)) \<cdot>\<^sub>B
(\<lbrace>Dom (fst f)\<rbrace> \<star>\<^sub>B \<lbrace>Dom (snd f)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)) =
\<lbrace>Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)\<rbrace>"
proof -
have 1: "E.Ide (Dom (fst f) \<^bold>\<star> Dom (snd f))"
- using f VV.arr_char arr_char dom_char
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))"
- using f VV.arr_char arr_char dom_char
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char dom_char
apply simp
by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1))
have 3: "\<^bold>\<lfloor>Dom (fst f) \<^bold>\<star> Dom (snd f)\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)\<^bold>\<rfloor>"
- using f VV.arr_char arr_char dom_char
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "(\<lbrace>Dom (fst f)\<rbrace> \<star>\<^sub>B \<lbrace>Dom (snd f)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)) =
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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
+ using 1 2 3 f VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.can_Ide_self B.vcomp_can by simp
qed
ultimately show ?thesis by simp
qed
finally show "dom (fst f \<star> snd f) = fst (VV.dom f) \<star> snd (VV.dom f)"
by simp
qed
show "\<And>f. VV.arr f \<Longrightarrow> cod (fst f \<star> snd f) = fst (VV.cod f) \<star> snd (VV.cod f)"
proof -
fix f
assume f: "VV.arr f"
have "cod (fst f \<star> snd f) = MkIde (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f))"
using f VV.arrE [of f] cod_char arr_hcomp hcomp_def by simp
also have "... = fst (VV.cod f) \<star> snd (VV.cod f)"
proof -
have "hcomp (fst (VV.cod f)) (snd (VV.cod f)) =
MkArr (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f))
(B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f)) \<cdot>\<^sub>B
(\<lbrace>Cod (fst f)\<rbrace> \<star>\<^sub>B \<lbrace>Cod (snd f)\<rbrace>) \<cdot>\<^sub>B
B.can (Cod (fst f) \<^bold>\<star> Cod (snd f)) (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f)) \<cdot>\<^sub>B
(\<lbrace>Cod (fst f)\<rbrace> \<star>\<^sub>B \<lbrace>Cod (snd f)\<rbrace>) \<cdot>\<^sub>B
B.can (Cod (fst f) \<^bold>\<star> Cod (snd f)) (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) =
\<lbrace>Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)\<rbrace>"
proof -
have 1: "E.Ide (Cod (fst f) \<^bold>\<star> Cod (snd f))"
- using f VV.arr_char arr_char dom_char
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f))"
- using f VV.arr_char arr_char dom_char
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char dom_char
apply simp
by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1))
have 3: "\<^bold>\<lfloor>Cod (fst f) \<^bold>\<star> Cod (snd f)\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)\<^bold>\<rfloor>"
- using f VV.arr_char arr_char dom_char
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "(\<lbrace>Cod (fst f)\<rbrace> \<star>\<^sub>B \<lbrace>Cod (snd f)\<rbrace>) \<cdot>\<^sub>B
B.can (Cod (fst f) \<^bold>\<star> Cod (snd f)) (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) =
B.can (Cod (fst f) \<^bold>\<star> Cod (snd f)) (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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
+ using 1 2 3 f VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.can_Ide_self B.vcomp_can by simp
qed
ultimately show ?thesis by simp
qed
finally show "cod (fst f \<star> snd f) = fst (VV.cod f) \<star> snd (VV.cod f)"
by simp
qed
show "\<And>g f. VV.seq g f \<Longrightarrow>
fst (VV.comp g f) \<star> snd (VV.comp g f) = (fst g \<star> snd g) \<cdot> (fst f \<star> snd f)"
proof -
fix f g
assume fg: "VV.seq g f"
have f: "arr (fst f) \<and> arr (snd f) \<and> src (fst f) = trg (snd f)"
- using fg VV.seq_char VV.arr_char by simp
+ using fg VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have g: "arr (fst g) \<and> arr (snd g) \<and> src (fst g) = trg (snd g)"
- using fg VV.seq_char VV.arr_char by simp
+ using fg VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have 1: "arr (fst (VV.comp g f)) \<and> arr (snd (VV.comp g f)) \<and>
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 \<cdot> fst f, snd g \<cdot> snd f)"
using fg 1 VV.comp_char VxV.comp_char
- by (metis (no_types, lifting) VV.seq_char VxV.seqE)
+ by (metis (no_types, lifting) VV.seq_char\<^sub>S\<^sub>b\<^sub>C VxV.seqE)
let ?X = "MkArr (Dom (fst (VV.comp g f)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd (VV.comp g f)))
(Cod (fst (VV.comp g f)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd (VV.comp g f)))
(B.can (Cod (fst (VV.comp g f)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd (VV.comp g f)))
(Cod (fst (VV.comp g f)) \<^bold>\<star> Cod (snd (VV.comp g f))) \<cdot>\<^sub>B
(Map (fst (VV.comp g f)) \<star>\<^sub>B Map (snd (VV.comp g f))) \<cdot>\<^sub>B
B.can (Dom (fst (VV.comp g f)) \<^bold>\<star> Dom (snd (VV.comp g f)))
(Dom (fst (VV.comp g f)) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd (VV.comp g f))))"
have 2: "fst (VV.comp g f) \<star> snd (VV.comp g f) = ?X"
unfolding hcomp_def using 1 by simp
also have "... = (fst g \<star> snd g) \<cdot> (fst f \<star> snd f)"
proof -
let ?GG = "MkArr (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)) (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g))
(B.can (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
(Map (fst g) \<star>\<^sub>B Map (snd g)) \<cdot>\<^sub>B
B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)))"
let ?FF = "MkArr (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)) (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f))
(B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f)) \<cdot>\<^sub>B
(Map (fst f) \<star>\<^sub>B Map (snd f)) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)))"
have 4: "arr ?FF \<and> arr ?GG \<and> Dom ?GG = Cod ?FF"
proof -
have "arr ?FF \<and> arr ?GG"
- using f g fg VV.arr_char VV.seqE hcomp_def A by presburger
+ using f g fg VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.seqE hcomp_def A by presburger
thus ?thesis
using 0 1 by (simp add: fg seq_char)
qed
have "(fst g \<star> snd g) \<cdot> (fst f \<star> snd f) = ?GG \<cdot> ?FF"
unfolding hcomp_def
- using 1 f g fg VV.arr_char VV.seqE by simp
+ using 1 f g fg VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<cdot> ?FF) = Dom ?X"
using fg 0 1 4 seq_char by simp
show "Cod (?GG \<cdot> ?FF) = Cod ?X"
using fg 0 1 4 seq_char by simp
show "Map (?GG \<cdot> ?FF) = Map ?X"
proof -
have "Map (?GG \<cdot> ?FF) = Map ?GG \<cdot>\<^sub>B Map ?FF"
using 4 by auto
also have
"... = (B.can (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
(Map (fst g) \<star>\<^sub>B Map (snd g)) \<cdot>\<^sub>B
B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g))) \<cdot>\<^sub>B
(B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f)) \<cdot>\<^sub>B
(Map (fst f) \<star>\<^sub>B Map (snd f)) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f)))"
using fg by simp
also have
"... = B.can (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
((Map (fst g) \<star>\<^sub>B Map (snd g)) \<cdot>\<^sub>B (Map (fst f) \<star>\<^sub>B Map (snd f))) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))"
proof -
have "(B.can (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
(Map (fst g) \<star>\<^sub>B Map (snd g)) \<cdot>\<^sub>B
B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g))) \<cdot>\<^sub>B
(B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f)) \<cdot>\<^sub>B
(Map (fst f) \<star>\<^sub>B Map (snd f)) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))) =
B.can (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
((Map (fst g) \<star>\<^sub>B Map (snd g)) \<cdot>\<^sub>B
(B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)) \<cdot>\<^sub>B
B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f))) \<cdot>\<^sub>B
(Map (fst f) \<star>\<^sub>B Map (snd f))) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))"
using B.comp_assoc by simp
also have "... = B.can (Cod (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
((Map (fst g) \<star>\<^sub>B Map (snd g)) \<cdot>\<^sub>B (Map (fst f) \<star>\<^sub>B Map (snd f))) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd f))"
proof -
have "(B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g))) \<cdot>\<^sub>B
(B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f))) =
\<lbrace>Cod (fst f) \<^bold>\<star> Cod (snd f)\<rbrace>"
proof -
have "(B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g))) \<cdot>\<^sub>B
(B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f))) =
B.can (Dom (fst g) \<^bold>\<star> Dom (snd g)) (Cod (fst f) \<^bold>\<star> Cod (snd f))"
proof -
have "E.Ide (Dom (fst g) \<^bold>\<star> Dom (snd g))"
using g arr_char
by (metis (no_types, lifting) E.Ide.simps(3) src_simps(2) trg_simps(2))
moreover have "E.Ide (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g))"
using 4 by auto
moreover have "E.Ide (Cod (fst f) \<^bold>\<star> Cod (snd f))"
using f arr_char
by (metis (no_types, lifting) E.Ide.simps(3) src_simps(2) trg_simps(2))
moreover have
"\<^bold>\<lfloor>Dom (fst g) \<^bold>\<star> Dom (snd g)\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)\<^bold>\<rfloor>"
using g E.Nml_HcompNml(1) calculation(1) by fastforce
moreover have
"\<^bold>\<lfloor>Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod (fst f) \<^bold>\<star> Cod (snd f)\<^bold>\<rfloor>"
using g fg seq_char
- by (metis (no_types, lifting) VV.seq_char VxV.seqE calculation(4))
+ by (metis (no_types, lifting) VV.seq_char\<^sub>S\<^sub>b\<^sub>C VxV.seqE calculation(4))
moreover have
"Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g) = Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)"
using 0 1 by (simp add: seq_char)
ultimately show ?thesis
using B.vcomp_can by simp
qed
also have "... = \<lbrace>Cod (fst f) \<^bold>\<star> Cod (snd f)\<rbrace>"
proof -
have "Dom (fst g) \<^bold>\<star> Dom (snd g) = Cod (fst f) \<^bold>\<star> Cod (snd f)"
- using 0 f g fg seq_char VV.seq_char VV.arr_char
+ using 0 f g fg seq_char VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
thus ?thesis
using f B.can_Ide_self [of "Dom (fst f) \<^bold>\<star> Dom (snd f)"]
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>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)) \<cdot>\<^sub>B
B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f))) \<cdot>\<^sub>B
(Map (fst f) \<star>\<^sub>B Map (snd f)) =
\<lbrace>Cod (fst f) \<^bold>\<star> Cod (snd f)\<rbrace> \<cdot>\<^sub>B (Map (fst f) \<star>\<^sub>B Map (snd f))"
by simp
also have "... = Map (fst f) \<star>\<^sub>B Map (snd f)"
proof -
have 1: "\<forall>p. arr p \<longrightarrow> map (cod p) \<cdot> map p = map p"
by blast
have 3: "\<lbrace>Cod (fst f)\<rbrace> \<cdot>\<^sub>B Map (fst f) = Map (map (cod (fst f)) \<cdot> map (fst f))"
by (simp add: f)
have 4: "map (cod (fst f)) \<cdot> map (fst f) = fst f"
using 1 f map_simp by simp
show ?thesis
proof -
have 2: "\<lbrace>Cod (snd f)\<rbrace> \<cdot>\<^sub>B Map (snd f) = Map (snd f)"
using 1 f map_simp
by (metis (no_types, lifting) Dom_cod Map_cod Map_comp arr_cod)
have "B.seq \<lbrace>Cod (snd f)\<rbrace> (Map (snd f))"
using f 2 by auto
moreover have "B.seq \<lbrace>Cod (fst f)\<rbrace> (Map (fst f))"
using 4 f 3 by auto
moreover have
"\<lbrace>Cod (fst f)\<rbrace> \<cdot>\<^sub>B Map (fst f) \<star>\<^sub>B \<lbrace>Cod (snd f)\<rbrace> \<cdot>\<^sub>B Map (snd f) =
Map (fst f) \<star>\<^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>\<star> Dom (snd g)) (Dom (fst g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd g)) \<cdot>\<^sub>B
B.can (Cod (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd f)) (Cod (fst f) \<^bold>\<star> Cod (snd f))) \<cdot>\<^sub>B
(Map (fst f) \<star>\<^sub>B Map (snd f)) =
Map (fst f) \<star>\<^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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd g)) (Cod (fst g) \<^bold>\<star> Cod (snd g)) \<cdot>\<^sub>B
(Map (fst g \<cdot> fst f) \<star>\<^sub>B Map (snd g \<cdot> snd f)) \<cdot>\<^sub>B
B.can (Dom (fst f) \<^bold>\<star> Dom (snd f)) (Dom (fst f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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
+ using 0 f g fg VV.seq_char\<^sub>S\<^sub>b\<^sub>C [of g f] VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char seq_char
by (metis (no_types, lifting) fst_conv)
hence "Map (fst g \<cdot> fst f) = Map (fst g) \<cdot>\<^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)) \<and>
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) \<star> snd (VV.comp g f) = (fst g \<star> snd g) \<cdot> (fst f \<star> 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 \<mu>" and "arr \<nu>" and "arr \<tau>"
and "src \<mu> = trg \<nu>" and "src \<nu> = trg \<tau>"
shows "(\<mu> \<star> \<nu>) \<star> \<tau> = \<mu> \<star> \<nu> \<star> \<tau>"
proof (intro arr_eqI)
have \<mu>\<nu>: "\<guillemotleft>Map \<mu> \<star>\<^sub>B Map \<nu> : \<lbrace>Dom \<mu> \<^bold>\<star> Dom \<nu>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod \<mu> \<^bold>\<star> Cod \<nu>\<rbrace>\<guillemotright>"
using assms src_def trg_def arr_char
by (auto simp add: E.eval_simps'(2-3) Pair_inject)
have \<nu>\<tau>: "\<guillemotleft>Map \<nu> \<star>\<^sub>B Map \<tau> : \<lbrace>Dom \<nu> \<^bold>\<star> Dom \<tau>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod \<nu> \<^bold>\<star> Cod \<tau>\<rbrace>\<guillemotright>"
using assms src_def trg_def arr_char
by (auto simp add: E.eval_simps'(2-3) Pair_inject)
show "hseq (\<mu> \<star> \<nu>) \<tau>"
using assms \<mu>\<nu> \<nu>\<tau> by auto
show "hseq \<mu> (\<nu> \<star> \<tau>)"
using assms \<mu>\<nu> \<nu>\<tau> by auto
show "Dom ((\<mu> \<star> \<nu>) \<star> \<tau>) = Dom (\<mu> \<star> \<nu> \<star> \<tau>)"
proof -
have "E.Nml (Dom \<mu>) \<and> E.Nml (Dom \<nu>) \<and> E.Nml (Dom \<tau>)"
using assms by blast
moreover have "E.Src (Dom \<mu>) = E.Trg (Dom \<nu>) \<and> E.Src (Dom \<nu>) = E.Trg (Dom \<tau>)"
using assms \<mu>\<nu> \<nu>\<tau>
by (metis (no_types, lifting) src_simps(2) trg_simps(2))
ultimately show ?thesis
using assms \<mu>\<nu> \<nu>\<tau> E.HcompNml_assoc by simp
qed
show "Cod ((\<mu> \<star> \<nu>) \<star> \<tau>) = Cod (\<mu> \<star> \<nu> \<star> \<tau>)"
proof -
have "E.Nml (Cod \<mu>) \<and> E.Nml (Cod \<nu>) \<and> E.Nml (Cod \<tau>)"
using assms by blast
moreover have "E.Src (Cod \<mu>) = E.Trg (Cod \<nu>) \<and> E.Src (Cod \<nu>) = E.Trg (Cod \<tau>)"
using assms \<mu>\<nu> \<nu>\<tau>
by (metis (no_types, lifting) arrE src_simps(2) trg_simps(2))
ultimately show ?thesis
using assms \<mu>\<nu> \<nu>\<tau> E.HcompNml_assoc by simp
qed
show "Map ((\<mu> \<star> \<nu>) \<star> \<tau>) = Map (\<mu> \<star> \<nu> \<star> \<tau>)"
proof -
have "Map ((\<mu> \<star> \<nu>) \<star> \<tau>) =
B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
proof -
have 1: "Map ((\<mu> \<star> \<nu>) \<star> \<tau>) =
B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
using assms \<mu>\<nu> \<nu>\<tau> 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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(B.can ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>)) \<cdot>\<^sub>B
B.can ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
proof -
have
"B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B Map \<tau> =
B.can ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>)"
proof -
have "B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>)
\<star>\<^sub>B Map \<tau> =
(B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<star>\<^sub>B B.can (Cod \<tau>) (Cod \<tau>)) \<cdot>\<^sub>B
((Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B Map \<tau>)"
proof -
have "B.seq (B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>))
((Map \<mu> \<star>\<^sub>B Map \<nu>) \<cdot>\<^sub>B B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>))"
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 \<tau>) (Cod \<tau>)) (Map \<tau>)"
using B.can_in_hom assms(3) by blast
moreover have "B.ide (B.can (Cod \<tau>) (Cod \<tau>))"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<star>\<^sub>B B.can (Cod \<tau>) (Cod \<tau>)) \<cdot>\<^sub>B
((Map \<mu> \<star>\<^sub>B Map \<nu>) \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B
B.can (Dom \<tau>) (Dom \<tau>))"
proof -
have "B.seq (Map \<mu> \<star>\<^sub>B Map \<nu>) (B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>))"
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 \<tau>) (B.can (Dom \<tau>) (Dom \<tau>))"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<star>\<^sub>B B.can (Cod \<tau>) (Cod \<tau>)) \<cdot>\<^sub>B
(B.can ((Cod \<mu> \<^bold>\<star> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<star> Dom \<nu>) \<^bold>\<star> Dom \<tau>)) \<cdot>\<^sub>B
(B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B
B.can (Dom \<tau>) (Dom \<tau>))"
proof -
have "(Map \<mu> \<star>\<^sub>B Map \<nu>) \<star>\<^sub>B Map \<tau> =
B.\<a>' \<lbrace>Cod \<mu>\<rbrace> \<lbrace>Cod \<nu>\<rbrace> \<lbrace>Cod \<tau>\<rbrace> \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
\<a>\<^sub>B \<lbrace>Dom \<mu>\<rbrace> \<lbrace>Dom \<nu>\<rbrace> \<lbrace>Dom \<tau>\<rbrace>"
using B.hcomp_reassoc(1)
by (metis (no_types, lifting) B.hcomp_in_vhomE B.in_homE \<mu>\<nu> \<nu>\<tau> arrE
assms(1) assms(2) assms(3))
also have "... = B.can ((Cod \<mu> \<^bold>\<star> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<star> Dom \<nu>) \<^bold>\<star> Dom \<tau>)"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<star>\<^sub>B B.can (Cod \<tau>) (Cod \<tau>)) \<cdot>\<^sub>B
(B.can ((Cod \<mu> \<^bold>\<star> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>))) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<star> Dom \<nu>) \<^bold>\<star> Dom \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B
B.can (Dom \<tau>) (Dom \<tau>)))"
using B.comp_assoc by simp
also have
"... = B.can ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>)"
proof -
have "(B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) (Cod \<mu> \<^bold>\<star> Cod \<nu>) \<star>\<^sub>B B.can (Cod \<tau>) (Cod \<tau>)) \<cdot>\<^sub>B
(B.can ((Cod \<mu> \<^bold>\<star> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>)) =
B.can ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>)"
proof -
have "E.Ide (Cod \<mu> \<^bold>\<star> Cod \<nu>)"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>)"
using E.Ide_HcompNml assms(1) assms(2) calculation by auto
moreover have "\<^bold>\<lfloor>Cod \<mu> \<^bold>\<star> Cod \<nu>\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>\<^bold>\<rfloor>"
using E.Nml_HcompNml(1) assms(1) assms(2) calculation(1) by fastforce
moreover have "E.Src (Cod \<mu> \<^bold>\<star> Cod \<nu>) = E.Trg (Cod \<tau>)"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) = E.Trg (Cod \<tau>)"
using E.Src_HcompNml assms(1) assms(2) calculation(1) calculation(4)
by fastforce
moreover have "\<^bold>\<lfloor>Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>\<^bold>\<rfloor> = \<^bold>\<lfloor>(Cod \<mu> \<^bold>\<star> Cod \<nu>) \<^bold>\<star> Cod \<tau>\<^bold>\<rfloor>"
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 \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<star> Dom \<nu>) \<^bold>\<star> Dom \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu> \<^bold>\<star> Dom \<nu>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<star>\<^sub>B B.can (Dom \<tau>) (Dom \<tau>)) =
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>)"
proof -
have "E.Ide (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>)"
by (metis (no_types, lifting) E.Ide_HcompNml arrE assms(1-2,4)
src_simps(2) trg_simps(2))
moreover have "E.Ide (Dom \<mu> \<^bold>\<star> Dom \<nu>)"
by (metis (no_types, lifting) E.Ide.simps(3) arrE assms(1-2,4)
src_simps(1) trg_simps(1))
moreover have "\<^bold>\<lfloor>Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom \<mu> \<^bold>\<star> Dom \<nu>\<^bold>\<rfloor>"
using E.Nml_HcompNml(1) assms(1-2) calculation(2) by fastforce
moreover have "E.Src (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) = E.Trg (Dom \<tau>)"
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 \<mu> \<^bold>\<star> Dom \<nu>) = E.Trg (Dom \<tau>)"
using E.Src_HcompNml assms(1-2) calculation(2) calculation(4)
by fastforce
moreover have "E.Ide ((Dom \<mu> \<^bold>\<star> Dom \<nu>) \<^bold>\<star> Dom \<tau>)"
using E.Ide.simps(3) assms(3) calculation(2) calculation(5) by blast
moreover have "\<^bold>\<lfloor>(Dom \<mu> \<^bold>\<star> Dom \<nu>) \<^bold>\<star> Dom \<tau>\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>\<^bold>\<rfloor>"
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 \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
B.can ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>)) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>) \<cdot>\<^sub>B
B.can ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
using B.comp_assoc by simp
also have "... = B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
proof -
have "B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
B.can ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) =
B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>)"
proof -
have "E.Ide (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>)"
using assms src_def trg_def by fastforce
moreover have "E.Ide ((Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>)"
using assms arr_char src_def trg_def E.Ide_HcompNml E.Src_HcompNml
by auto
moreover have "E.Ide (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>)"
using assms arr_char src_def trg_def
by (simp add: E.Nml_HcompNml(1) E.Ide_HcompNml E.Trg_HcompNml)
moreover have "\<^bold>\<lfloor>Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>\<^bold>\<rfloor> = \<^bold>\<lfloor>(Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>\<^bold>\<rfloor>"
using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp
moreover have "\<^bold>\<lfloor>(Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu>) \<^bold>\<star> Cod \<tau>\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>\<^bold>\<rfloor>"
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 \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>) \<cdot>\<^sub>B
B.can ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>) =
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
proof -
have "E.Ide (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>)"
using assms src_def trg_def by fastforce
moreover have "E.Ide ((Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>)"
using assms arr_char src_def trg_def E.Ide_HcompNml E.Src_HcompNml
by auto
moreover have "E.Ide (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
using assms arr_char src_def trg_def
by (simp add: E.Nml_HcompNml(1) E.Ide_HcompNml E.Trg_HcompNml)
moreover have "\<^bold>\<lfloor>Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>\<^bold>\<rfloor> = \<^bold>\<lfloor>(Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>\<^bold>\<rfloor>"
using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp
moreover have
"\<^bold>\<lfloor>(Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu>) \<^bold>\<star> Dom \<tau>\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>\<^bold>\<rfloor>"
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 (\<mu> \<star> \<nu> \<star> \<tau>)"
proof -
have 1: "Map (\<mu> \<star> \<nu> \<star> \<tau>) =
B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
using assms Map_hcomp [of \<mu> "\<nu> \<star> \<tau>"] Map_hcomp [of \<nu> \<tau>] by simp
also have
"... = B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) \<cdot>\<^sub>B
((B.can (Cod \<mu>) (Cod \<mu>) \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>)) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu>) (Dom \<mu>) \<star>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>))) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
proof -
have "Map \<mu> \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>) =
(B.can (Cod \<mu>) (Cod \<mu>) \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>)) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B (Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>))"
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 \<mu>) (Cod \<mu>) \<star>\<^sub>B
B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>)) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu>) (Dom \<mu>) \<star>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>))"
using assms B.interchange B.comp_arr_dom [of "Map \<mu>" "B.can (Dom \<mu>) (Dom \<mu>)"]
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 \<mu> \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>) =
(B.can (Cod \<mu>) (Cod \<mu>) \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>)) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
(B.can (Dom \<mu>) (Dom \<mu>) \<star>\<^sub>B B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>))"
by simp
thus ?thesis by simp
qed
also have
"... = (B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) \<cdot>\<^sub>B
(B.can (Cod \<mu>) (Cod \<mu>) \<star>\<^sub>B B.can (Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<nu> \<^bold>\<star> Cod \<tau>))) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
((B.can (Dom \<mu>) (Dom \<mu>) \<star>\<^sub>B
B.can (Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>))"
using B.comp_assoc by simp
also have "... = B.can (Cod \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod \<tau>) (Cod \<mu> \<^bold>\<star> Cod \<nu> \<^bold>\<star> Cod \<tau>) \<cdot>\<^sub>B
(Map \<mu> \<star>\<^sub>B Map \<nu> \<star>\<^sub>B Map \<tau>) \<cdot>\<^sub>B
B.can (Dom \<mu> \<^bold>\<star> Dom \<nu> \<^bold>\<star> Dom \<tau>) (Dom \<mu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<nu> \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom \<tau>)"
using assms \<mu>\<nu> \<nu>\<tau> 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 \<longleftrightarrow> endo a \<and> E.Obj (Dom a) \<and> Map a = \<lbrace>Dom a\<rbrace>"
proof
assume a: "obj a"
show "endo a \<and> E.Obj (Dom a) \<and> Map a = \<lbrace>Dom a\<rbrace>"
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 = \<lbrace>Dom a\<rbrace>"
using a ide_char src_def by blast
qed
next
assume a: "endo a \<and> E.Obj (Dom a) \<and> Map a = \<lbrace>Dom a\<rbrace>"
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 \<star> 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 \<star> 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 \<star> 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 \<star> a) = Map a"
using "1" Map_ide(1) assms by fastforce
qed
lemma hcomp_ide_src:
assumes "ide f"
shows "f \<star> 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 \<star> 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 \<star> 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 \<star> src f) = Map f"
by (simp add: "1" Map_ide(1) assms)
qed
lemma hcomp_trg_ide:
assumes "ide f"
shows "trg f \<star> 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 \<star> 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 \<star> 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 \<star> 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)) \<star> MkArr (Dom a') (Cod a) (Map g)"
using f by auto
also have "... = MkIde (E.Trg (Cod a)) \<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom a') (E.Trg (Cod a) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod a)
(B.can (E.Trg (Cod a) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod a) (E.Trg (Cod a) \<^bold>\<star> Cod a) \<cdot>\<^sub>B
(\<lbrace>E.Trg (Cod a)\<rbrace> \<star>\<^sub>B Map g) \<cdot>\<^sub>B
B.can (E.Trg (Cod a) \<^bold>\<star> Dom a') (E.Trg (Cod a) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> Cod a) \<cdot>\<^sub>B
(trg\<^sub>B \<lbrace>Cod a\<rbrace> \<star>\<^sub>B Map g) \<cdot>\<^sub>B
B.can (E.Trg (Cod a) \<^bold>\<star> Dom a') (Dom a'))"
proof -
have "E.Trg (Cod a) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod a = Cod a"
using a a' arr_char E.HcompNml_Trg_Nml by blast
moreover have "\<lbrace>E.Trg (Cod a)\<rbrace> = trg\<^sub>B \<lbrace>Cod a\<rbrace>"
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 \<lbrace>Cod a\<rbrace> \<cdot>\<^sub>B (trg\<^sub>B \<lbrace>Cod a\<rbrace> \<star>\<^sub>B Map g) \<cdot>\<^sub>B B.lunit' \<lbrace>Dom a'\<rbrace>)"
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
\<open>Cod g = Dom a\<close> \<open>Dom g = Dom a'\<close>
by (metis (no_types, lifting) Cod.simps(1) in_homE)
moreover have "B.can (Cod a) (E.Trg (Cod a) \<^bold>\<star> Cod a) = B.lunit \<lbrace>Cod a\<rbrace>"
using a ide_char arr_char B.canE_unitor(2) by blast
moreover have "B.can (E.Trg (Dom a') \<^bold>\<star> Dom a') (Dom a') = B.lunit' \<lbrace>Dom a'\<rbrace>"
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 \<lbrace>Cod a\<rbrace> = 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 \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B (trg\<^sub>B (Map g) \<star>\<^sub>B Map g) \<cdot>\<^sub>B B.lunit' \<lbrace>Dom g\<rbrace> = Map g"
proof -
have "B.lunit \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B (trg\<^sub>B (Map g) \<star>\<^sub>B Map g) \<cdot>\<^sub>B B.lunit' \<lbrace>Dom g\<rbrace> =
B.lunit \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B B.lunit' \<lbrace>Cod g\<rbrace> \<cdot>\<^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 \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B B.lunit' \<lbrace>Cod g\<rbrace>) \<cdot>\<^sub>B Map g"
using B.comp_assoc by simp
also have "... = \<lbrace>Cod g\<rbrace> \<cdot>\<^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 \<lbrace>Cod a\<rbrace> \<cdot>\<^sub>B (trg\<^sub>B \<lbrace>Cod a\<rbrace> \<star>\<^sub>B Map g) \<cdot>\<^sub>B B.lunit' \<lbrace>Dom a'\<rbrace> = 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 "\<exists>f. in_hom f a' a \<and> 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) \<star> src (MkArr (Dom a') (Cod a) (Map g))"
using f by auto
also have "... = MkArr (Dom a') (Cod a) (Map g) \<star> MkIde (E.Src (Cod a))"
using a a' f src_def [of a] vconn_implies_hpar by auto
also have "... = MkArr (Dom a' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Cod a)) (Cod a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Cod a))
(B.can (Cod a \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Cod a)) (Cod a \<^bold>\<star> E.Src (Cod a)) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B \<lbrace>E.Src (Cod a)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom a' \<^bold>\<star> E.Src (Cod a)) (Dom a' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> E.Src (Cod a)) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B src\<^sub>B \<lbrace>Cod a\<rbrace>) \<cdot>\<^sub>B
B.can (Dom a' \<^bold>\<star> E.Src (Cod a)) (Dom a'))"
proof -
have "Dom a' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Cod a) = Cod a"
using a a' arr_char E.HcompNml_Nml_Src by blast
moreover have "\<lbrace>E.Src (Cod a)\<rbrace> = src\<^sub>B \<lbrace>Cod a\<rbrace>"
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 \<lbrace>Cod a\<rbrace> \<cdot>\<^sub>B (Map g \<star>\<^sub>B src\<^sub>B \<lbrace>Cod a\<rbrace>) \<cdot>\<^sub>B B.runit' \<lbrace>Dom a'\<rbrace>)"
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 \<lbrace>Cod a\<rbrace> = 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 \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B (Map g \<star>\<^sub>B src\<^sub>B (Map g)) \<cdot>\<^sub>B B.runit' \<lbrace>Dom g\<rbrace> = Map g"
proof -
have "B.runit \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B (Map g \<star>\<^sub>B src\<^sub>B (Map g)) \<cdot>\<^sub>B B.runit' \<lbrace>Dom g\<rbrace> =
B.runit \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B B.runit'\<lbrace>Cod g\<rbrace> \<cdot>\<^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 \<lbrace>Cod g\<rbrace> \<cdot>\<^sub>B B.runit' \<lbrace>Cod g\<rbrace>) \<cdot>\<^sub>B Map g"
using B.comp_assoc by simp
also have "... = \<lbrace>Cod g\<rbrace> \<cdot>\<^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 \<lbrace>Cod a\<rbrace> \<cdot>\<^sub>B (Map g \<star>\<^sub>B src\<^sub>B \<lbrace>Cod a\<rbrace>) \<cdot>\<^sub>B B.runit' \<lbrace>Dom a'\<rbrace> = 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 "\<exists>f. in_hom f a' a \<and> 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' \<and> 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) \<star>\<^sub>B Map f"
using par by auto
also have "... = trg\<^sub>B (Map f') \<star>\<^sub>B Map f'"
proof -
have "\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B Map f = \<lbrace>E.Trg (Dom f')\<rbrace> \<star>\<^sub>B Map f'"
proof -
have A: "\<guillemotleft>B.can (E.Trg (Dom f) \<^bold>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f) :
\<lbrace>E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B \<lbrace>Dom f\<rbrace>\<guillemotright>"
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: "\<guillemotleft>B.can (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f) (E.Trg (Dom f) \<^bold>\<star> Cod f) :
\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B \<lbrace>Cod f\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f\<rbrace>\<guillemotright>"
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: "\<guillemotleft>\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B Map f :
\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B \<lbrace>Dom f\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B \<lbrace>Cod f\<rbrace>\<guillemotright>"
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: "(\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B Map f) \<cdot>\<^sub>B
B.can (E.Trg (Dom f) \<^bold>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f) =
(\<lbrace>E.Trg (Dom f')\<rbrace> \<star>\<^sub>B Map f') \<cdot>\<^sub>B
B.can (E.Trg (Dom f') \<^bold>\<star> Dom f') (E.Trg (Dom f') \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f')"
proof -
have 2: "B.can (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f) (E.Trg (Dom f) \<^bold>\<star> Cod f) \<cdot>\<^sub>B
(\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B Map f) \<cdot>\<^sub>B
B.can (E.Trg (Dom f) \<^bold>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f) =
B.can (E.Trg (Dom f') \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f') (E.Trg (Dom f') \<^bold>\<star> Cod f') \<cdot>\<^sub>B
(\<lbrace>E.Trg (Dom f')\<rbrace> \<star>\<^sub>B Map f') \<cdot>\<^sub>B
B.can (E.Trg (Dom f') \<^bold>\<star> Dom f') (E.Trg (Dom f') \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f) (E.Trg (Dom f) \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f) (E.Trg (Dom f) \<^bold>\<star> Cod f))
((\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B Map f) \<cdot>\<^sub>B
B.can (E.Trg (Dom f) \<^bold>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f))"
using A B C by auto
moreover have
"B.seq (B.can (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f) (E.Trg (Dom f) \<^bold>\<star> Cod f))
((\<lbrace>E.Trg (Dom f')\<rbrace> \<star>\<^sub>B Map f') \<cdot>\<^sub>B
B.can (E.Trg (Dom f') \<^bold>\<star> Dom f') (E.Trg (Dom f') \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f'))"
using par 1 2 arr_char calculation(2) by auto
moreover have "B.can (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f) (E.Trg (Dom f) \<^bold>\<star> Cod f) =
B.can (E.Trg (Dom f') \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod f') (E.Trg (Dom f') \<^bold>\<star> 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>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 (\<lbrace>E.Trg (Dom f)\<rbrace> \<star>\<^sub>B Map f)
(B.can (E.Trg (Dom f) \<^bold>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom f))"
using A C by auto
moreover have "B.seq (\<lbrace>E.Trg (Dom f')\<rbrace> \<star>\<^sub>B Map f')
(B.can (E.Trg (Dom f) \<^bold>\<star> Dom f) (E.Trg (Dom f) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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) = \<lbrace>E.Trg (Dom f)\<rbrace> \<and>
trg\<^sub>B (Map f') = \<lbrace>E.Trg (Dom f')\<rbrace>"
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' \<and> 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 \<star>\<^sub>B src\<^sub>B (Map f)"
using par apply simp by (metis B.hseqE B.hseq_char')
also have "... = Map f' \<star>\<^sub>B src\<^sub>B (Map f')"
proof -
have "Map f \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace> = Map f' \<star>\<^sub>B \<lbrace>E.Src (Dom f')\<rbrace>"
proof -
have 2: "E.Ide (Cod f \<^bold>\<star> E.Src (Dom f))"
using par arr_char src.preserves_arr by auto
hence 3: "E.Ide (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f))"
using par arr_char E.Nml_Src E.Ide_HcompNml calculation by auto
have 4: "\<^bold>\<lfloor>Cod f \<^bold>\<star> E.Src (Dom f)\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)\<^bold>\<rfloor>"
using par arr_char by (simp add: E.Nml_HcompNml(1))
have A: "\<guillemotleft>B.can (Dom f \<^bold>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) :
\<lbrace>Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Dom f\<rbrace> \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace>\<guillemotright>"
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: "\<guillemotleft>B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) (Cod f \<^bold>\<star> E.Src (Dom f)) :
\<lbrace>Cod f\<rbrace> \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)\<rbrace>\<guillemotright>"
using 2 3 4 B.can_in_hom [of "Cod f \<^bold>\<star> E.Src (Dom f)" "Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)"]
by simp
have C: "\<guillemotleft>Map f \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace> :
\<lbrace>Dom f\<rbrace> \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod f\<rbrace> \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace>\<guillemotright>"
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 \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) =
(Map f' \<star>\<^sub>B \<lbrace>E.Src (Dom f')\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f' \<^bold>\<star> E.Src (Dom f')) (Dom f' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f'))"
proof -
have 6: "B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) (Cod f \<^bold>\<star> E.Src (Dom f)) \<cdot>\<^sub>B
(Map f \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) =
B.can (Cod f' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f')) (Cod f' \<^bold>\<star> E.Src (Dom f')) \<cdot>\<^sub>B
(Map f' \<star>\<^sub>B \<lbrace>E.Src (Dom f')\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f' \<^bold>\<star> E.Src (Dom f')) (Dom f' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) (Cod f \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) (Cod f \<^bold>\<star> E.Src (Dom f)))
((Map f \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)))"
using A B C by auto
moreover have
"B.seq (B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) (Cod f \<^bold>\<star> E.Src (Dom f)))
((Map f' \<star>\<^sub>B \<lbrace>E.Src (Dom f')\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f' \<^bold>\<star> E.Src (Dom f')) (Dom f' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f')))"
using par 1 6 arr_char calculation(2) by auto
moreover have "B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)) (Cod f \<^bold>\<star> E.Src (Dom f)) =
B.can (Cod f' \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f')) (Cod f' \<^bold>\<star> 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>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 \<star>\<^sub>B \<lbrace>E.Src (Dom f)\<rbrace>)
(B.can (Dom f \<^bold>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> E.Src (Dom f)))"
using A C by auto
moreover have "B.seq (Map f' \<star>\<^sub>B \<lbrace>E.Src (Dom f')\<rbrace>)
(B.can (Dom f \<^bold>\<star> E.Src (Dom f)) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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) = \<lbrace>E.Src (Dom f)\<rbrace> \<and>
src\<^sub>B (Map f') = \<lbrace>E.Src (Dom f')\<rbrace>"
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 \<a>
where "\<a> \<tau> \<mu> \<nu> \<equiv> if VVV.arr (\<tau>, \<mu>, \<nu>) then hcomp \<tau> (hcomp \<mu> \<nu>) else null"
interpretation natural_isomorphism VVV.comp vcomp HoHV HoVH
\<open>\<lambda>\<tau>\<mu>\<nu>. \<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))\<close>
proof
show "\<And>\<tau>\<mu>\<nu>. \<not> VVV.arr \<tau>\<mu>\<nu> \<Longrightarrow> \<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>)) = null"
using \<a>_def by simp
show "\<And>\<tau>\<mu>\<nu>. VVV.arr \<tau>\<mu>\<nu> \<Longrightarrow>
dom (\<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))) = HoHV (VVV.dom \<tau>\<mu>\<nu>)"
- using VVV.arr_char VV.arr_char \<a>_def hcomp_assoc HoHV_def VVV.dom_simp VV.dom_simp
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<a>_def hcomp_assoc HoHV_def VVV.dom_simp VV.dom_simp
by force
show 1: "\<And>\<tau>\<mu>\<nu>. VVV.arr \<tau>\<mu>\<nu> \<Longrightarrow>
cod (\<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))) = HoVH (VVV.cod \<tau>\<mu>\<nu>)"
- using VVV.arr_char VV.arr_char \<a>_def HoVH_def VVV.cod_simp VV.cod_simp by force
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C \<a>_def HoVH_def VVV.cod_simp VV.cod_simp by force
show "\<And>\<tau>\<mu>\<nu>. VVV.arr \<tau>\<mu>\<nu> \<Longrightarrow>
HoVH \<tau>\<mu>\<nu> \<cdot>
\<a> (fst (VVV.dom \<tau>\<mu>\<nu>)) (fst (snd (VVV.dom \<tau>\<mu>\<nu>)))
(snd (snd (VVV.dom \<tau>\<mu>\<nu>))) =
\<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))"
using \<a>_def HoVH.as_nat_trans.is_natural_1 HoVH_def by auto
show "\<And>\<tau>\<mu>\<nu>. VVV.arr \<tau>\<mu>\<nu> \<Longrightarrow>
\<a> (fst (VVV.cod \<tau>\<mu>\<nu>)) (fst (snd (VVV.cod \<tau>\<mu>\<nu>)))
(snd (snd (VVV.cod \<tau>\<mu>\<nu>))) \<cdot> HoHV \<tau>\<mu>\<nu> =
\<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))"
proof -
fix \<tau>\<mu>\<nu>
assume \<tau>\<mu>\<nu>: "VVV.arr \<tau>\<mu>\<nu>"
have "HoHV \<tau>\<mu>\<nu> = \<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))"
unfolding \<a>_def HoHV_def
- using \<tau>\<mu>\<nu> HoHV.preserves_cod hcomp_assoc VVV.arr_char VV.arr_char
+ using \<tau>\<mu>\<nu> HoHV.preserves_cod hcomp_assoc VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
thus "\<a> (fst (VVV.cod \<tau>\<mu>\<nu>)) (fst (snd (VVV.cod \<tau>\<mu>\<nu>))) (snd (snd (VVV.cod \<tau>\<mu>\<nu>))) \<cdot>
HoHV \<tau>\<mu>\<nu> =
\<a> (fst \<tau>\<mu>\<nu>) (fst (snd \<tau>\<mu>\<nu>)) (snd (snd \<tau>\<mu>\<nu>))"
using 1 \<tau>\<mu>\<nu> comp_cod_arr \<a>_def
by (metis (no_types, lifting) HoVH_def HoHV.preserves_arr prod.collapse)
qed
show "\<And>fgh. VVV.ide fgh \<Longrightarrow> iso (\<a> (fst fgh) (fst (snd fgh)) (snd (snd fgh)))"
using \<a>_def HoVH.preserves_ide HoVH_def by auto
qed
definition \<i>
where "\<i> \<equiv> \<lambda>a. a"
sublocale bicategory vcomp hcomp \<a> \<i> src trg
- using hcomp_obj_self \<a>_def hcomp_assoc VVV.arr_char VV.arr_char
+ using hcomp_obj_self \<a>_def hcomp_assoc VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply unfold_locales
by (auto simp add: \<i>_def ide_in_hom(2))
lemma is_bicategory:
shows "bicategory vcomp hcomp \<a> \<i> src trg"
..
sublocale strict_bicategory vcomp hcomp \<a> \<i> src trg
proof
show "\<And>fgh. ide fgh \<Longrightarrow> 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 "\<guillemotleft>fgh : trg fgh \<star> fgh \<Rightarrow> fgh\<guillemotright>"
using fgh hcomp_def hcomp_trg_ide by auto
show "trg fgh \<star> fgh = (\<i> (trg fgh) \<star> fgh) \<cdot> \<a>' (trg fgh) (trg fgh) fgh"
proof -
have "(\<i> (trg fgh) \<star> fgh) \<cdot> \<a>' (trg fgh) (trg fgh) fgh =
(trg fgh \<star> fgh) \<cdot> \<a>' (trg fgh) (trg fgh) fgh"
using fgh \<i>_def by metis
also have "... = (trg fgh \<star> fgh) \<cdot> (trg fgh \<star> trg fgh \<star> fgh)"
using fgh \<a>_def by fastforce
also have "... = trg fgh \<star> 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 "\<And>fgh. ide fgh \<Longrightarrow> 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 "\<guillemotleft>fgh : fgh \<star> src fgh \<Rightarrow> fgh\<guillemotright>"
using fgh hcomp_def hcomp_ide_src by auto
show "fgh \<star> src fgh = (fgh \<star> \<i> (src fgh)) \<cdot> \<a> fgh (src fgh) (src fgh)"
proof -
have "(fgh \<star> \<i> (src fgh)) \<cdot> \<a> fgh (src fgh) (src fgh) =
(fgh \<star> src fgh) \<cdot> \<a> fgh (src fgh) (src fgh)"
using fgh \<i>_def by metis
also have "... = (fgh \<star> src fgh) \<cdot> (fgh \<star> src fgh \<star> src fgh)"
using fgh \<a>_def by fastforce
also have "... = fgh \<star> 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 "\<And>f g h. \<lbrakk> ide f; ide g; ide h; src f = trg g; src g = trg h \<rbrakk> \<Longrightarrow> ide (\<a> f g h)"
- using \<a>_def VV.arr_char VVV.arr_char by auto
+ using \<a>_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
theorem is_strict_bicategory:
shows "strict_bicategory vcomp hcomp \<a> \<i> src trg"
..
lemma iso_char:
shows "iso \<mu> \<longleftrightarrow> arr \<mu> \<and> B.iso (Map \<mu>)"
and "iso \<mu> \<Longrightarrow> inv \<mu> = MkArr (Cod \<mu>) (Dom \<mu>) (B.inv (Map \<mu>))"
proof -
have 1: "iso \<mu> \<Longrightarrow> arr \<mu> \<and> B.iso (Map \<mu>)"
proof -
assume \<mu>: "iso \<mu>"
obtain \<nu> where \<nu>: "inverse_arrows \<mu> \<nu>"
using \<mu> by auto
have "B.inverse_arrows (Map \<mu>) (Map \<nu>)"
proof
show "B.ide (Map \<mu> \<cdot>\<^sub>B Map \<nu>)"
proof -
have "Map \<mu> \<cdot>\<^sub>B Map \<nu> = Map (\<mu> \<cdot> \<nu>)"
using \<mu> \<nu> inverse_arrows_def Map_comp arr_char seq_char
by (metis (no_types, lifting) ide_compE)
moreover have "B.ide ..."
using \<nu> ide_char by blast
ultimately show ?thesis by simp
qed
show "B.ide (Map \<nu> \<cdot>\<^sub>B Map \<mu>)"
proof -
have "Map \<nu> \<cdot>\<^sub>B Map \<mu> = Map (\<nu> \<cdot> \<mu>)"
using \<mu> \<nu> inverse_arrows_def comp_char [of \<nu> \<mu>] by simp
moreover have "B.ide ..."
using \<nu> ide_char by blast
ultimately show ?thesis by simp
qed
qed
thus "arr \<mu> \<and> B.iso (Map \<mu>)"
using \<mu> by auto
qed
let ?\<nu> = "MkArr (Cod \<mu>) (Dom \<mu>) (B.inv (Map \<mu>))"
have 2: "arr \<mu> \<and> B.iso (Map \<mu>) \<Longrightarrow> iso \<mu> \<and> inv \<mu> = ?\<nu>"
proof
assume \<mu>: "arr \<mu> \<and> B.iso (Map \<mu>)"
have \<nu>: "\<guillemotleft>?\<nu> : cod \<mu> \<Rightarrow> dom \<mu>\<guillemotright>"
using \<mu> arr_char dom_char cod_char by auto
have 4: "inverse_arrows \<mu> ?\<nu>"
proof
show "ide (?\<nu> \<cdot> \<mu>)"
proof -
have "?\<nu> \<cdot> \<mu> = dom \<mu>"
using \<mu> \<nu> MkArr_Map comp_char seq_char B.comp_inv_arr' dom_char by auto
thus ?thesis
using \<mu> by simp
qed
show "ide (\<mu> \<cdot> ?\<nu>)"
proof -
have "\<mu> \<cdot> ?\<nu> = cod \<mu>"
using \<mu> \<nu> MkArr_Map comp_char seq_char B.comp_arr_inv' cod_char by auto
thus ?thesis
using \<mu> by simp
qed
qed
thus "iso \<mu>" by auto
show "inv \<mu> = ?\<nu>"
using 4 inverse_unique by simp
qed
have 3: "arr \<mu> \<and> B.iso (Map \<mu>) \<Longrightarrow> iso \<mu>"
using 2 by simp
show "iso \<mu> \<longleftrightarrow> arr \<mu> \<and> B.iso (Map \<mu>)"
using 1 3 by blast
show "iso \<mu> \<Longrightarrow> inv \<mu> = ?\<nu>"
using 1 2 by blast
qed
subsection "The Strictness Theorem"
text \<open>
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.
\<close>
text \<open>
We define a map \<open>UP\<close> from the given bicategory \<open>B\<close> 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.
\<close>
definition UP\<^sub>0
where "UP\<^sub>0 a \<equiv> if B.obj a then MkIde \<^bold>\<langle>a\<^bold>\<rangle>\<^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>\<langle>a\<^bold>\<rangle>\<^sub>0"] src_def by simp
lemma UP\<^sub>0_in_hom [intro]:
assumes "B.obj a"
shows "\<guillemotleft>UP\<^sub>0 a : UP\<^sub>0 a \<rightarrow> UP\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>UP\<^sub>0 a : UP\<^sub>0 a \<Rightarrow> UP\<^sub>0 a\<guillemotright>"
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 \<mu> \<equiv> if B.arr \<mu> then MkArr \<^bold>\<langle>B.dom \<mu>\<^bold>\<rangle> \<^bold>\<langle>B.cod \<mu>\<^bold>\<rangle> \<mu> else null"
lemma Dom_UP [simp]:
assumes "B.arr \<mu>"
shows "Dom (UP \<mu>) = \<^bold>\<langle>B.dom \<mu>\<^bold>\<rangle>"
using assms UP_def by fastforce
lemma Cod_UP [simp]:
assumes "B.arr \<mu>"
shows "Cod (UP \<mu>) = \<^bold>\<langle>B.cod \<mu>\<^bold>\<rangle>"
using assms UP_def by fastforce
lemma Map_UP [simp]:
assumes "B.arr \<mu>"
shows "Map (UP \<mu>) = \<mu>"
using assms UP_def by fastforce
lemma arr_UP:
assumes "B.arr \<mu>"
shows "arr (UP \<mu>)"
using assms UP_def
by (intro arrI, fastforce+)
lemma UP_in_hom [intro]:
assumes "B.arr \<mu>"
shows "\<guillemotleft>UP \<mu> : UP\<^sub>0 (src\<^sub>B \<mu>) \<rightarrow> UP\<^sub>0 (trg\<^sub>B \<mu>)\<guillemotright>"
and "\<guillemotleft>UP \<mu> : UP (B.dom \<mu>) \<Rightarrow> UP (B.cod \<mu>)\<guillemotright>"
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 \<mu>"
shows "arr (UP \<mu>)"
and "src (UP \<mu>) = UP\<^sub>0 (src\<^sub>B \<mu>)" and "trg (UP \<mu>) = UP\<^sub>0 (trg\<^sub>B \<mu>)"
and "dom (UP \<mu>) = UP (B.dom \<mu>)" and "cod (UP \<mu>) = UP (B.cod \<mu>)"
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 \<mu>
assume \<mu>: "B.arr \<mu>"
show "isomorphic (UP (src\<^sub>B \<mu>)) (src (UP \<mu>))"
proof -
let ?\<phi> = "MkArr \<^bold>\<langle>src\<^sub>B \<mu>\<^bold>\<rangle> \<^bold>\<langle>src\<^sub>B \<mu>\<^bold>\<rangle>\<^sub>0 (src\<^sub>B \<mu>)"
have \<phi>: "\<guillemotleft>?\<phi> : UP (src\<^sub>B \<mu>) \<Rightarrow> src (UP \<mu>)\<guillemotright>"
proof
show 1: "arr ?\<phi>"
using \<mu> by (intro arrI, auto)
show "dom ?\<phi> = UP (src\<^sub>B \<mu>)"
using \<mu> 1 dom_char UP_def by simp
show "cod ?\<phi> = src (UP \<mu>)"
using \<mu> 1 cod_char src_def by auto
qed
have "iso ?\<phi>"
using \<mu> \<phi> iso_char src_def by auto
thus ?thesis
using \<phi> isomorphic_def by auto
qed
show "isomorphic (UP (trg\<^sub>B \<mu>)) (trg (UP \<mu>))"
proof -
let ?\<phi> = "MkArr \<^bold>\<langle>trg\<^sub>B \<mu>\<^bold>\<rangle> \<^bold>\<langle>trg\<^sub>B \<mu>\<^bold>\<rangle>\<^sub>0 (trg\<^sub>B \<mu>)"
have \<phi>: "\<guillemotleft>?\<phi> : UP (trg\<^sub>B \<mu>) \<Rightarrow> trg (UP \<mu>)\<guillemotright>"
proof
show 1: "arr ?\<phi>"
using \<mu> by (intro arrI, auto)
show "dom ?\<phi> = UP (trg\<^sub>B \<mu>)"
using \<mu> 1 dom_char UP_def by simp
show "cod ?\<phi> = trg (UP \<mu>)"
using \<mu> 1 cod_char trg_def by auto
qed
have "iso ?\<phi>"
using \<mu> \<phi> iso_char trg_def by auto
thus ?thesis
using \<phi> isomorphic_def by auto
qed
qed
interpretation HoUP_UP: composite_functor B.VV.comp VV.comp vcomp
UP.FF \<open>\<lambda>\<mu>\<nu>. hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> ..
interpretation UPoH: composite_functor B.VV.comp V\<^sub>B vcomp
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close> UP ..
abbreviation \<Phi>\<^sub>o
where "\<Phi>\<^sub>o fg \<equiv> MkArr (\<^bold>\<langle>fst fg\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>snd fg\<^bold>\<rangle>) \<^bold>\<langle>fst fg \<star>\<^sub>B snd fg\<^bold>\<rangle> (fst fg \<star>\<^sub>B snd fg)"
interpretation \<Phi>: transformation_by_components
B.VV.comp vcomp HoUP_UP.map UPoH.map \<Phi>\<^sub>o
proof
fix fg
assume fg: "B.VV.ide fg"
show "\<guillemotleft>\<Phi>\<^sub>o fg : HoUP_UP.map fg \<Rightarrow> UPoH.map fg\<guillemotright>"
proof (intro in_homI)
show 1: "arr (\<Phi>\<^sub>o fg)"
- using fg arr_char B.VV.ide_char B.VV.arr_char by auto
+ using fg arr_char B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
show "dom (\<Phi>\<^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
+ using 1 fg UP.FF_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C dom_char hcomp_def B.can_Ide_self
by simp
show "cod (\<Phi>\<^sub>o fg) = UPoH.map fg"
- using fg arr_char cod_char B.VV.ide_char B.VV.arr_char UP_def by auto
+ using fg arr_char cod_char B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C UP_def by auto
qed
next
fix \<mu>\<nu>
assume \<mu>\<nu>: "B.VV.arr \<mu>\<nu>"
show "\<Phi>\<^sub>o (B.VV.cod \<mu>\<nu>) \<cdot> HoUP_UP.map \<mu>\<nu> = UPoH.map \<mu>\<nu> \<cdot> \<Phi>\<^sub>o (B.VV.dom \<mu>\<nu>)"
proof -
have "\<Phi>\<^sub>o (B.VV.cod \<mu>\<nu>) \<cdot> HoUP_UP.map \<mu>\<nu> =
MkArr (\<^bold>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
(fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)"
proof -
have "\<Phi>\<^sub>o (B.VV.cod \<mu>\<nu>) \<cdot> HoUP_UP.map \<mu>\<nu> =
MkArr (\<^bold>\<langle>B.cod (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>) (\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
(B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)) \<cdot>
MkArr (\<^bold>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
(fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)"
- using \<mu>\<nu> B.VV.arr_char arr_char UP.FF_def hcomp_def UP_def B.VV.cod_simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
((B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)) \<cdot>\<^sub>B (fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>))"
- using \<mu>\<nu> B.VV.arr_char
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by (intro comp_MkArr arr_MkArr, auto)
also have "... = MkArr (\<^bold>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
(fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)"
- using \<mu>\<nu> B.VV.arr_char B.comp_cod_arr by auto
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.comp_cod_arr by auto
finally show ?thesis by simp
qed
also have "... = UPoH.map \<mu>\<nu> \<cdot> \<Phi>\<^sub>o (B.VV.dom \<mu>\<nu>)"
proof -
have "UPoH.map \<mu>\<nu> \<cdot> \<Phi>\<^sub>o (B.VV.dom \<mu>\<nu>) =
MkArr (\<^bold>\<langle>B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
(fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>) \<cdot>
MkArr (\<^bold>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>))"
- using \<mu>\<nu> B.VV.arr_char arr_char UP.FF_def hcomp_def UP_def B.VV.dom_simp
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
((fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>) \<cdot>\<^sub>B (B.dom (fst \<mu>\<nu>) \<star>\<^sub>B B.dom (snd \<mu>\<nu>)))"
- using \<mu>\<nu> B.VV.arr_char arr_MkArr
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_MkArr
apply (intro comp_MkArr arr_MkArr) by auto
also have "... = MkArr (\<^bold>\<langle>B.dom (fst \<mu>\<nu>)\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>B.dom (snd \<mu>\<nu>)\<^bold>\<rangle>)
(\<^bold>\<langle>B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)\<^bold>\<rangle>)
(fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)"
- using \<mu>\<nu> B.VV.arr_char B.comp_arr_dom by auto
+ using \<mu>\<nu> B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<equiv> \<Phi>.map"
lemma cmp\<^sub>U\<^sub>P_in_hom [intro]:
assumes "B.arr (fst \<mu>\<nu>)" and "B.arr (snd \<mu>\<nu>)" and "src\<^sub>B (fst \<mu>\<nu>) = trg\<^sub>B (snd \<mu>\<nu>)"
shows "\<guillemotleft>cmp\<^sub>U\<^sub>P \<mu>\<nu> : UP\<^sub>0 (src\<^sub>B (snd \<mu>\<nu>)) \<rightarrow> UP\<^sub>0 (trg\<^sub>B (fst \<mu>\<nu>))\<guillemotright>"
and "\<guillemotleft>cmp\<^sub>U\<^sub>P \<mu>\<nu> : UP (B.dom (fst \<mu>\<nu>)) \<star> UP (B.dom (snd \<mu>\<nu>))
\<Rightarrow> UP (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))\<guillemotright>"
proof -
let ?\<mu> = "fst \<mu>\<nu>" and ?\<nu> = "snd \<mu>\<nu>"
show 1: "\<guillemotleft>cmp\<^sub>U\<^sub>P \<mu>\<nu> :
UP (B.dom ?\<mu>) \<star> UP (B.dom ?\<nu>) \<Rightarrow> UP (B.cod ?\<mu> \<star>\<^sub>B B.cod ?\<nu>)\<guillemotright>"
proof
show "arr (cmp\<^sub>U\<^sub>P \<mu>\<nu>)"
using assms by auto
show "dom (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = UP (B.dom ?\<mu>) \<star> UP (B.dom ?\<nu>)"
proof -
have "B.VV.in_hom (?\<mu>, ?\<nu>) (B.dom ?\<mu>, B.dom ?\<nu>) (B.cod ?\<mu>, B.cod ?\<nu>)"
- using assms B.VV.in_hom_char B.VV.arr_char by auto
+ using assms B.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
hence "dom (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = HoUP_UP.map (B.dom ?\<mu>, B.dom ?\<nu>)"
by auto
also have "... = UP (B.dom ?\<mu>) \<star> UP (B.dom ?\<nu>)"
using assms UP.FF_def by fastforce
finally show ?thesis by simp
qed
show "cod (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = UP (B.cod ?\<mu> \<star>\<^sub>B B.cod ?\<nu>)"
- using assms B.VV.in_hom_char B.VV.arr_char B.VV.cod_simp by auto
+ using assms B.VV.in_hom_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.cod_simp by auto
qed
show "\<guillemotleft>cmp\<^sub>U\<^sub>P \<mu>\<nu> : UP\<^sub>0 (src\<^sub>B ?\<nu>) \<rightarrow> UP\<^sub>0 (trg\<^sub>B ?\<mu>)\<guillemotright>"
using assms 1 src_dom [of "cmp\<^sub>U\<^sub>P \<mu>\<nu>"] trg_dom [of "cmp\<^sub>U\<^sub>P \<mu>\<nu>"] by fastforce
qed
lemma cmp\<^sub>U\<^sub>P_simps [simp]:
assumes "B.arr (fst \<mu>\<nu>)" and "B.arr (snd \<mu>\<nu>)" and "src\<^sub>B (fst \<mu>\<nu>) = trg\<^sub>B (snd \<mu>\<nu>)"
shows "arr (cmp\<^sub>U\<^sub>P \<mu>\<nu>)"
and "src (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = UP\<^sub>0 (src\<^sub>B (snd \<mu>\<nu>))" and "trg (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = UP\<^sub>0 (trg\<^sub>B (fst \<mu>\<nu>))"
and "dom (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = UP (B.dom (fst \<mu>\<nu>)) \<star> UP (B.dom (snd \<mu>\<nu>))"
and "cod (cmp\<^sub>U\<^sub>P \<mu>\<nu>) = UP (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))"
using assms cmp\<^sub>U\<^sub>P_in_hom [of \<mu>\<nu>] 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>\<langle>fst fg\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>snd fg\<^bold>\<rangle>"
and "Cod (cmp\<^sub>U\<^sub>P fg) = \<^bold>\<langle>fst fg \<star>\<^sub>B snd fg\<^bold>\<rangle>"
and "Map (cmp\<^sub>U\<^sub>P fg) = fst fg \<star>\<^sub>B snd fg"
- using assms B.VV.ide_char B.VV.arr_char by auto
+ using assms B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
interpretation \<Phi>: 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 \<Phi>.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>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g)"
- using assms B.VV.ide_char B.VV.arr_char by simp
+ using assms B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (f \<star>\<^sub>B g)"
using assms cmp\<^sub>U\<^sub>P_ide_simp iso_char \<Phi>.components_are_iso [of "(f, g)"]
- B.VV.ide_char B.VV.arr_char
+ B.VV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
interpretation UP: pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \<a> \<i> 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 \<a>\<^sub>B[f, g, h] \<cdot> cmp\<^sub>U\<^sub>P (f \<star>\<^sub>B g, h) \<cdot> (cmp\<^sub>U\<^sub>P (f, g) \<star> UP h) =
cmp\<^sub>U\<^sub>P (f, g \<star>\<^sub>B h) \<cdot> (UP f \<star> cmp\<^sub>U\<^sub>P (g, h)) \<cdot> \<a> (UP f) (UP g) (UP h)"
proof -
have "UP \<a>\<^sub>B[f, g, h] \<cdot> cmp\<^sub>U\<^sub>P (f \<star>\<^sub>B g, h) \<cdot> (cmp\<^sub>U\<^sub>P (f, g) \<star> UP h) =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> (f \<star>\<^sub>B g \<star>\<^sub>B h)"
proof -
have 1: "UP.hseq\<^sub>D (MkIde \<^bold>\<langle>f\<^bold>\<rangle>) (MkIde \<^bold>\<langle>g\<^bold>\<rangle>)"
using f g fg hseq_char src_def trg_def arr_char by auto
have 2: "UP.hseq\<^sub>D (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g) \<cdot> MkIde (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>))
(MkIde \<^bold>\<langle>h\<^bold>\<rangle>)"
proof -
have "MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g) \<cdot> MkIde (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g)"
proof -
have "MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g) \<cdot> MkIde (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g) \<cdot> MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (f \<star>\<^sub>B g)"
using f g fg by simp
also have "... = MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> ((f \<star>\<^sub>B g) \<cdot>\<^sub>B (f \<star>\<^sub>B g))"
using f g fg by (intro comp_MkArr arr_MkArr, auto)
also have "... = MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^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 \<a>\<^sub>B[f, g, h] = MkArr \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle> \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> \<a>\<^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
+ using f g h fg gh UP_def B.HoHV_def B.HoVH_def B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have
"cmp\<^sub>U\<^sub>P (f \<star>\<^sub>B g, h) = MkArr (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle> ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) ((f \<star>\<^sub>B g) \<star>\<^sub>B h)"
- using f g h fg gh \<Phi>.map_simp_ide \<Phi>.map_def B.VV.arr_char UP.FF_def B.VV.cod_simp
+ using f g h fg gh \<Phi>.map_simp_ide \<Phi>.map_def B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C UP.FF_def B.VV.cod_simp
hcomp_def B.can_Ide_self
by simp
moreover have "cmp\<^sub>U\<^sub>P (f, g) \<star> UP h =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (B.inv \<a>\<^sub>B[f, g, h])"
proof -
have "MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>)
(B.can (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<cdot>\<^sub>B (f \<star>\<^sub>B g) \<cdot>\<^sub>B B.can (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>)) =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (f \<star>\<^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>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^sub>B g) \<cdot>
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) (f \<star>\<^sub>B g) =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> (f \<star>\<^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>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) = B.inv \<a>\<^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 \<Phi>.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
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.can_Ide_self B.comp_cod_arr B.VV.cod_simp
by simp
qed
ultimately have "UP \<a>\<^sub>B[f, g, h] \<cdot> cmp\<^sub>U\<^sub>P (f \<star>\<^sub>B g, h) \<cdot> (cmp\<^sub>U\<^sub>P (f, g) \<star> UP h) =
MkArr \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle> \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> \<a>\<^sub>B[f, g, h] \<cdot>
MkArr (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle> ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (B.inv \<a>\<^sub>B[f, g, h])"
using comp_assoc by presburger
also have "... = MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle>
(\<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B
B.inv \<a>\<^sub>B[f, g, h])"
proof -
have "Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (B.inv \<a>\<^sub>B[f, g, h])) \<and>
Arr (MkArr (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) ((f \<star>\<^sub>B g) \<star>\<^sub>B h)) \<and>
Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>)
(((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B B.inv \<a>\<^sub>B[f, g, h])) \<and>
Arr (MkArr (\<^bold>\<langle>f \<star>\<^sub>B g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle> ((f \<star>\<^sub>B g) \<star>\<^sub>B h)) \<and>
Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle>
(((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B B.inv \<a>\<^sub>B[f, g, h])) \<and>
Arr (MkArr \<^bold>\<langle>(f \<star>\<^sub>B g) \<star>\<^sub>B h\<^bold>\<rangle> \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> \<a>\<^sub>B[f, g, h])"
using f g h fg gh B.\<alpha>.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>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> (f \<star>\<^sub>B g \<star>\<^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 \<star>\<^sub>B h) \<cdot> (UP f \<star> cmp\<^sub>U\<^sub>P (g, h)) \<cdot> \<a> (UP f) (UP g) (UP h)"
proof -
have "cmp\<^sub>U\<^sub>P (f, g \<star>\<^sub>B h) \<cdot> (UP f \<star> cmp\<^sub>U\<^sub>P (g, h)) \<cdot> \<a> (UP f) (UP g) (UP h) =
cmp\<^sub>U\<^sub>P (f, g \<star>\<^sub>B h) \<cdot> (MkIde \<^bold>\<langle>f\<^bold>\<rangle> \<star> cmp\<^sub>U\<^sub>P (g, h)) \<cdot>
(MkIde \<^bold>\<langle>f\<^bold>\<rangle> \<star> MkIde \<^bold>\<langle>g\<^bold>\<rangle> \<star> MkIde \<^bold>\<langle>h\<^bold>\<rangle>)"
- using f g h fg gh VVV.arr_char VV.arr_char arr_char src_def trg_def UP_def \<a>_def
+ using f g h fg gh VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char src_def trg_def UP_def \<a>_def
by auto
also have "... = MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> (f \<star>\<^sub>B g \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h)"
proof -
have "cmp\<^sub>U\<^sub>P (f, g \<star>\<^sub>B h) = MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> (f \<star>\<^sub>B g \<star>\<^sub>B h) \<cdot>
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h)"
using f g h fg gh \<Phi>.map_simp_ide \<Phi>.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
+ B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle> (g \<star>\<^sub>B h)"
using g h gh cmp\<^sub>U\<^sub>P_ide_simp by blast
moreover have "MkIde \<^bold>\<langle>f\<^bold>\<rangle> \<star> MkArr (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle> (g \<star>\<^sub>B h) =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^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>\<langle>f\<^bold>\<rangle> \<star> MkIde \<^bold>\<langle>g\<^bold>\<rangle> \<star> MkIde \<^bold>\<langle>h\<^bold>\<rangle> =
MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h)"
proof -
have "\<guillemotleft>f : f \<Rightarrow>\<^sub>B f\<guillemotright> \<and> \<guillemotleft>g : g \<Rightarrow>\<^sub>B g\<guillemotright> \<and> \<guillemotleft>h : h \<Rightarrow>\<^sub>B h\<guillemotright>"
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>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> (f \<star>\<^sub>B g \<star>\<^sub>B h)"
proof -
have "Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<and>
Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<and>
Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<and>
Arr (MkArr (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g \<star>\<^sub>B h\<^bold>\<rangle>) \<^bold>\<langle>f \<star>\<^sub>B g \<star>\<^sub>B h\<^bold>\<rangle> (f \<star>\<^sub>B g \<star>\<^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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \<a> \<i> 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 \<mu> f g
assume f: "B.ide f" and g: "B.ide g"
assume \<mu>: "\<guillemotleft>\<mu> : UP f \<Rightarrow> UP g\<guillemotright>"
show "\<exists>\<nu>. \<guillemotleft>\<nu> : f \<Rightarrow>\<^sub>B g\<guillemotright> \<and> UP \<nu> = \<mu>"
proof -
have 1: "\<guillemotleft>Map \<mu> : f \<Rightarrow>\<^sub>B g\<guillemotright>"
using f g \<mu> UP_def arr_char in_hom_char by auto
moreover have "UP (Map \<mu>) = \<mu>"
proof -
have "\<mu> = MkArr (Dom \<mu>) (Cod \<mu>) (Map \<mu>)"
using \<mu> MkArr_Map by auto
also have "... = UP (Map \<mu>)"
using "1" B.arrI UP.as_nat_trans.preserves_hom UP_def \<mu> 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 ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>") (* Inherited from functor, I think. *)
lemma Map_reflects_hhom:
assumes "B.obj a" and "B.obj b" and "ide g"
and "\<guillemotleft>g : UP.map\<^sub>0 a \<rightarrow> UP.map\<^sub>0 b\<guillemotright>"
shows "\<guillemotleft>Map g : a \<rightarrow>\<^sub>B b\<guillemotright>"
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 "\<lbrace>Dom g\<rbrace> = 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: "... = \<lbrace>E.Src (Dom a)\<rbrace>"
using assms src_def by auto
also have "... = \<lbrace>\<^bold>\<langle>Map a\<^bold>\<rangle>\<^sub>0\<rbrace>"
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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \<a> \<i> src trg UP cmp\<^sub>U\<^sub>P
proof
(* UP is full, hence locally full. *)
show "\<And>f f' \<nu>. \<lbrakk> B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f';
\<guillemotleft>\<nu> : UP f \<Rightarrow> UP f'\<guillemotright> \<rbrakk> \<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright> \<and> UP \<mu> = \<nu>"
using UP.is_full by simp
(* UP is biessentially surjective on objects. *)
show "\<And>b. obj b \<Longrightarrow> \<exists>a. B.obj a \<and> 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>\<langle>Map b\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>Map b\<^bold>\<rangle>\<^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 ?\<phi> = "MkArr \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0 (Dom b) (Map b)"
have \<phi>: "arr ?\<phi>"
proof -
have 2: "E.Obj (Dom b)"
using b obj_char by blast
have "E.Nml \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0 \<and> E.Ide \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0"
using 1 by auto
moreover have "E.Nml (Dom b) \<and> E.Ide (Dom b)"
using b arr_char [of b] by auto
moreover have "E.Src \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0 = E.Src (Dom b)"
using b 1 2
apply (cases "Dom b")
apply simp_all
by fastforce
moreover have "E.Trg \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0 = E.Trg (Dom b)"
using b 1 2
apply (cases "Dom b")
apply simp_all
by fastforce
moreover have "\<guillemotleft>Map b : \<lbrace>\<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Dom b\<rbrace>\<guillemotright>"
using b 1 by (elim objE, auto)
ultimately show ?thesis
using arr_char \<open>E.Nml \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0 \<and> E.Ide \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0\<close> by auto
qed
hence "iso ?\<phi>"
using 1 iso_char by auto
moreover have "dom ?\<phi> = UP.map\<^sub>0 (Map b)"
using \<phi> dom_char b 1 3 B.objE UP.map\<^sub>0_def UP_def src_def by auto
moreover have "cod ?\<phi> = b"
using \<phi> cod_char b 4 1 by auto
ultimately have "isomorphic (UP.map\<^sub>0 (Map b)) b"
using \<phi> 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 "\<exists>a. B.obj a \<and> equivalent_objects (UP.map\<^sub>0 a) b"
using 1 6 by auto
qed
(* UP is locally essentially surjective. *)
show "\<And>a b g. \<lbrakk> B.obj a; B.obj b; \<guillemotleft>g : UP.map\<^sub>0 a \<rightarrow> UP.map\<^sub>0 b\<guillemotright>; ide g \<rbrakk> \<Longrightarrow>
\<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.ide f \<and> isomorphic (UP f) g"
proof -
fix a b g
assume a: "B.obj a" and b: "B.obj b"
assume g_in_hhom: "\<guillemotleft>g : UP.map\<^sub>0 a \<rightarrow> UP.map\<^sub>0 b\<guillemotright>"
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 = \<lbrace>Dom g\<rbrace>"
using ide_g by simp
have Map_g_in_hhom: "\<guillemotleft>Map g : a \<rightarrow>\<^sub>B b\<guillemotright>"
using a b ide_g g_in_hhom Map_reflects_hhom by simp
let ?\<phi> = "MkArr \<^bold>\<langle>Map g\<^bold>\<rangle> (Dom g) (Map g)"
have \<phi>: "arr ?\<phi>"
proof -
have "\<guillemotleft>Map ?\<phi> : \<lbrace>Dom ?\<phi>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod ?\<phi>\<rbrace>\<guillemotright>"
using 1 Map_g_eq by auto
moreover have "E.Ide \<^bold>\<langle>Map g\<^bold>\<rangle> \<and> E.Nml \<^bold>\<langle>Map g\<^bold>\<rangle>"
using 1 by simp
moreover have "E.Ide (Dom g) \<and> E.Nml (Dom g)"
using ide_g arr_char ide_char by blast
moreover have "E.Src \<^bold>\<langle>Map g\<^bold>\<rangle> = 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 \<open>arr (UP a)\<close> a in_hhomE UP_def)
moreover have "E.Trg \<^bold>\<langle>Map g\<^bold>\<rangle> = E.Trg (Dom g)"
proof -
have "E.Trg \<^bold>\<langle>Map g\<^bold>\<rangle> = \<^bold>\<langle>b\<^bold>\<rangle>\<^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>\<langle>b\<^bold>\<rangle>\<^sub>0"
using b \<open>arr (UP b)\<close> 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 "\<guillemotleft>?\<phi> : UP (Map g) \<Rightarrow> g\<guillemotright>"
by (simp add: "1" \<phi> ide_g in_hom_char)
moreover have "iso ?\<phi>"
using \<phi> 1 iso_char by simp
ultimately have "isomorphic (UP (Map g)) g"
using isomorphic_def by auto
thus "\<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.ide f \<and> 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 \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \<a> \<i> src trg
UP cmp\<^sub>U\<^sub>P"
..
text \<open>
Next, we work out the details of the equivalence pseudofunctor \<open>DN\<close> in the
converse direction.
\<close>
definition DN
where "DN \<mu> \<equiv> if arr \<mu> then Map \<mu> else B.null"
lemma DN_in_hom [intro]:
assumes "arr \<mu>"
shows "\<guillemotleft>DN \<mu> : DN (src \<mu>) \<rightarrow>\<^sub>B DN (trg \<mu>)\<guillemotright>"
and "\<guillemotleft>DN \<mu> : DN (dom \<mu>) \<Rightarrow>\<^sub>B DN (cod \<mu>)\<guillemotright>"
using assms DN_def arr_char [of \<mu>] B.vconn_implies_hpar(1-2) E.eval_in_hom(1)
B.in_hhom_def
by auto
lemma DN_simps [simp]:
assumes "arr \<mu>"
shows "B.arr (DN \<mu>)"
and "src\<^sub>B (DN \<mu>) = DN (src \<mu>)" and "trg\<^sub>B (DN \<mu>) = DN (trg \<mu>)"
and "B.dom (DN \<mu>) = DN (dom \<mu>)" and "B.cod (DN \<mu>) = DN (cod \<mu>)"
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 \<mu>
assume \<mu>: "arr \<mu>"
show "B.isomorphic (DN (src \<mu>)) (src\<^sub>B (DN \<mu>))"
proof -
have "DN (src \<mu>) = src\<^sub>B (DN \<mu>)"
using B.src.is_extensional DN_def DN_simps(2) by auto
moreover have "B.ide (DN (src \<mu>))"
using \<mu> by simp
ultimately show ?thesis
using \<mu> B.isomorphic_reflexive by auto
qed
show "B.isomorphic (DN (trg \<mu>)) (trg\<^sub>B (DN \<mu>))"
proof -
have "DN (trg \<mu>) = trg\<^sub>B (DN \<mu>)"
using \<open>B.isomorphic (DN (src \<mu>)) (src\<^sub>B (DN \<mu>))\<close> by fastforce
moreover have "B.ide (DN (trg \<mu>))"
using \<mu> 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 \<open>\<lambda>\<mu>\<nu>. H\<^sub>B (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> ..
interpretation DNoH: composite_functor VV.comp vcomp V\<^sub>B
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> DN ..
abbreviation \<Psi>\<^sub>o
where "\<Psi>\<^sub>o fg \<equiv> B.can (Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)) (Dom (fst fg) \<^bold>\<star> Dom (snd fg))"
abbreviation \<Psi>\<^sub>o'
where "\<Psi>\<^sub>o' fg \<equiv> B.can (Dom (fst fg) \<^bold>\<star> Dom (snd fg)) (Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg))"
lemma \<Psi>\<^sub>o_in_hom:
assumes "VV.ide fg"
shows "\<guillemotleft>\<Psi>\<^sub>o fg : Map (fst fg) \<star>\<^sub>B Map (snd fg) \<Rightarrow>\<^sub>B \<lbrace>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<rbrace>\<guillemotright>"
and "\<guillemotleft>\<Psi>\<^sub>o' fg : \<lbrace>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<rbrace> \<Rightarrow>\<^sub>B Map (fst fg) \<star>\<^sub>B Map (snd fg)\<guillemotright>"
and "B.inverse_arrows (\<Psi>\<^sub>o fg) (\<Psi>\<^sub>o' fg)"
proof -
have 1: "E.Ide (Dom (fst fg) \<^bold>\<star> Dom (snd fg))"
unfolding E.Ide.simps(3)
apply (intro conjI)
- using assms VV.ide_char VV.arr_char arr_char
+ using assms VV.ide_char VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char
apply simp
- using VV.arr_char VV.ideD(1) assms
+ using VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg))"
using 1
- by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char assms)
+ by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ideD(1) arr_char assms)
have 3: "\<^bold>\<lfloor>Dom (fst fg) \<^bold>\<star> Dom (snd fg)\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<^bold>\<rfloor>"
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)
+ E.Nmlize_Nml VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ideD(1) arr_char assms 1)
have 4: "\<lbrace>Dom (fst fg) \<^bold>\<star> Dom (snd fg)\<rbrace> = Map (fst fg) \<star>\<^sub>B Map (snd fg)"
- using assms VV.ide_char VV.arr_char arr_char by simp
+ using assms VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by simp
show "\<guillemotleft>\<Psi>\<^sub>o fg : Map (fst fg) \<star>\<^sub>B Map (snd fg) \<Rightarrow>\<^sub>B \<lbrace>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<rbrace>\<guillemotright>"
using 1 2 3 4 by auto
show "\<guillemotleft>\<Psi>\<^sub>o' fg : \<lbrace>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<rbrace> \<Rightarrow>\<^sub>B Map (fst fg) \<star>\<^sub>B Map (snd fg)\<guillemotright>"
using 1 2 3 4 by auto
show "B.inverse_arrows (\<Psi>\<^sub>o fg) (\<Psi>\<^sub>o' fg)"
using 1 2 3 B.inverse_arrows_can by blast
qed
interpretation \<Psi>: transformation_by_components
VV.comp V\<^sub>B HoDN_DN.map DNoH.map \<Psi>\<^sub>o
proof
fix fg
assume fg: "VV.ide fg"
have 1: "\<lbrace>Dom (fst fg) \<^bold>\<star> Dom (snd fg)\<rbrace> = Map (fst fg) \<star>\<^sub>B Map (snd fg)"
- using fg VV.ide_char VV.arr_char arr_char by simp
+ using fg VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char by simp
show "\<guillemotleft>\<Psi>\<^sub>o fg : HoDN_DN.map fg \<Rightarrow>\<^sub>B DNoH.map fg\<guillemotright>"
proof
show "B.arr (\<Psi>\<^sub>o fg)"
using fg \<Psi>\<^sub>o_in_hom by blast
show "B.dom (\<Psi>\<^sub>o fg) = HoDN_DN.map fg"
proof -
have "B.dom (\<Psi>\<^sub>o fg) = Map (fst fg) \<star>\<^sub>B Map (snd fg)"
using fg \<Psi>\<^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
+ using fg DN.FF_def DN_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def VV.ide_char\<^sub>S\<^sub>b\<^sub>C by auto
finally show ?thesis by simp
qed
show "B.cod (\<Psi>\<^sub>o fg) = DNoH.map fg"
proof -
have "B.cod (\<Psi>\<^sub>o fg) = \<lbrace>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<rbrace>"
using fg \<Psi>\<^sub>o_in_hom by blast
also have "... = DNoH.map fg"
proof -
have "DNoH.map fg =
B.can (Cod (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd fg)) (Cod (fst fg) \<^bold>\<star> Cod (snd fg)) \<cdot>\<^sub>B
(Map (fst fg) \<star>\<^sub>B Map (snd fg)) \<cdot>\<^sub>B
B.can (Dom (fst fg) \<^bold>\<star> Dom (snd fg)) (Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg))"
- using fg DN_def Map_hcomp VV.arr_char
+ using fg DN_def Map_hcomp VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply simp
using VV.ideD(1) by blast
also have "... =
B.can (Cod (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd fg)) (Cod (fst fg) \<^bold>\<star> Cod (snd fg)) \<cdot>\<^sub>B
B.can (Dom (fst fg) \<^bold>\<star> Dom (snd fg)) (Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg))"
proof -
have "(Map (fst fg) \<star>\<^sub>B Map (snd fg)) \<cdot>\<^sub>B
B.can (Dom (fst fg) \<^bold>\<star> Dom (snd fg)) (Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)) =
B.can (Dom (fst fg) \<^bold>\<star> Dom (snd fg)) (Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg))"
using fg 1 \<Psi>\<^sub>o_in_hom B.comp_cod_arr by blast
thus ?thesis by simp
qed
also have "... = \<lbrace>Dom (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd fg)\<rbrace>"
proof -
have "B.can (Cod (fst fg) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd fg)) (Cod (fst fg) \<^bold>\<star> Cod (snd fg)) = \<Psi>\<^sub>o fg"
- using fg VV.ide_char Cod_ide by simp
+ using fg VV.ide_char\<^sub>S\<^sub>b\<^sub>C Cod_ide by simp
thus ?thesis
using fg 1 \<Psi>\<^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 "\<And>f. VV.arr f \<Longrightarrow>
\<Psi>\<^sub>o (VV.cod f) \<cdot>\<^sub>B HoDN_DN.map f = DNoH.map f \<cdot>\<^sub>B \<Psi>\<^sub>o (VV.dom f)"
proof -
fix \<mu>\<nu>
assume \<mu>\<nu>: "VV.arr \<mu>\<nu>"
show "\<Psi>\<^sub>o (VV.cod \<mu>\<nu>) \<cdot>\<^sub>B HoDN_DN.map \<mu>\<nu> = DNoH.map \<mu>\<nu> \<cdot>\<^sub>B \<Psi>\<^sub>o (VV.dom \<mu>\<nu>)"
proof -
have 1: "E.Ide (Dom (fst \<mu>\<nu>) \<^bold>\<star> Dom (snd \<mu>\<nu>))"
unfolding E.Ide.simps(3)
by (metis (no_types, lifting) VV.arrE \<mu>\<nu> arrE src_simps(2) trg_simps(2))
have 2: "E.Ide (Dom (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd \<mu>\<nu>))"
using 1
- by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char \<mu>\<nu>)
+ by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ideD(1) arr_char \<mu>\<nu>)
have 3: "\<^bold>\<lfloor>Dom (fst \<mu>\<nu>) \<^bold>\<star> Dom (snd \<mu>\<nu>)\<^bold>\<rfloor> = \<^bold>\<lfloor>Dom (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd \<mu>\<nu>)\<^bold>\<rfloor>"
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 \<mu>\<nu> 1)
+ E.Nmlize_Nml VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char \<mu>\<nu> 1)
have 4: "E.Ide (Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>))"
unfolding E.Ide.simps(3)
by (metis (no_types, lifting) "1" E.Ide.simps(3) VV.arrE \<mu>\<nu> arr_char)
have 5: "E.Ide (Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>))"
using 4
- by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char \<mu>\<nu>)
+ by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.ideD(1) arr_char \<mu>\<nu>)
have 6: "\<^bold>\<lfloor>Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)\<^bold>\<rfloor> = \<^bold>\<lfloor>Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>)\<^bold>\<rfloor>"
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 \<mu>\<nu> 1)
+ E.Nmlize_Nml VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char \<mu>\<nu> 1)
have A: "\<guillemotleft>\<Psi>\<^sub>o' \<mu>\<nu> : \<lbrace>Dom (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom (snd \<mu>\<nu>)\<rbrace>
\<Rightarrow>\<^sub>B \<lbrace>Dom (fst \<mu>\<nu>) \<^bold>\<star> Dom (snd \<mu>\<nu>)\<rbrace>\<guillemotright>"
using 1 2 3 by auto
have B: "\<guillemotleft>Map (fst \<mu>\<nu>) \<star>\<^sub>B Map (snd \<mu>\<nu>) :
\<lbrace>Dom (fst \<mu>\<nu>) \<^bold>\<star> Dom (snd \<mu>\<nu>)\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)\<rbrace>\<guillemotright>"
- using \<mu>\<nu> VV.arr_char arr_char src_def trg_def E.Nml_implies_Arr E.eval_simps'(2-3)
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char src_def trg_def E.Nml_implies_Arr E.eval_simps'(2-3)
by auto
have C: "\<guillemotleft>B.can (Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>)) (Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)) :
\<lbrace>Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>)\<rbrace>\<guillemotright>"
using 4 5 6 by auto
have "\<Psi>\<^sub>o (VV.cod \<mu>\<nu>) \<cdot>\<^sub>B HoDN_DN.map \<mu>\<nu> =
B.can (Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>)) (Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)) \<cdot>\<^sub>B
(Map (fst \<mu>\<nu>) \<star>\<^sub>B Map (snd \<mu>\<nu>))"
- using \<mu>\<nu> VV.arr_char VV.cod_char arr_char src_def trg_def cod_char DN.FF_def DN_def
+ using \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C arr_char src_def trg_def cod_char DN.FF_def DN_def
by auto
also have "... = B.can (Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>))
(Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)) \<cdot>\<^sub>B
(Map (fst \<mu>\<nu>) \<star>\<^sub>B Map (snd \<mu>\<nu>)) \<cdot>\<^sub>B \<Psi>\<^sub>o' \<mu>\<nu> \<cdot>\<^sub>B \<Psi>\<^sub>o \<mu>\<nu>"
- using B \<mu>\<nu> VV.arr_char arr_char src_def trg_def
+ using B \<mu>\<nu> VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<mu>\<nu> \<cdot>\<^sub>B \<Psi>\<^sub>o (VV.dom \<mu>\<nu>)"
proof -
have "DNoH.map \<mu>\<nu> \<cdot>\<^sub>B \<Psi>\<^sub>o (VV.dom \<mu>\<nu>) =
B.can (Cod (fst \<mu>\<nu>) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod (snd \<mu>\<nu>)) (Cod (fst \<mu>\<nu>) \<^bold>\<star> Cod (snd \<mu>\<nu>)) \<cdot>\<^sub>B
(Map (fst \<mu>\<nu>) \<star>\<^sub>B Map (snd \<mu>\<nu>)) \<cdot>\<^sub>B \<Psi>\<^sub>o' \<mu>\<nu> \<cdot>\<^sub>B \<Psi>\<^sub>o (VV.dom \<mu>\<nu>)"
- using \<mu>\<nu> DN_def VV.arr_char B.comp_assoc by simp
+ using \<mu>\<nu> DN_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.comp_assoc by simp
moreover have "\<Psi>\<^sub>o (VV.dom \<mu>\<nu>) = \<Psi>\<^sub>o \<mu>\<nu>"
- using \<mu>\<nu> VV.dom_char VV.arr_char by auto
+ using \<mu>\<nu> VV.dom_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<equiv> \<Psi>.map"
interpretation \<Psi>: natural_isomorphism VV.comp V\<^sub>B HoDN_DN.map DNoH.map cmp\<^sub>D\<^sub>N
using \<Psi>\<^sub>o_in_hom B.iso_def \<Psi>.map_simp_ide
apply unfold_locales
apply auto
by blast
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
lemma cmp\<^sub>D\<^sub>N_in_hom [intro]:
assumes "arr (fst \<mu>\<nu>)" and "arr (snd \<mu>\<nu>)" and "src (fst \<mu>\<nu>) = trg (snd \<mu>\<nu>)"
shows "\<guillemotleft>cmp\<^sub>D\<^sub>N \<mu>\<nu> : DN (src (snd \<mu>\<nu>)) \<rightarrow>\<^sub>B DN (trg (fst \<mu>\<nu>))\<guillemotright>"
and "\<guillemotleft>cmp\<^sub>D\<^sub>N \<mu>\<nu> : DN (dom (fst \<mu>\<nu>)) \<star>\<^sub>B DN (dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B DN (cod (fst \<mu>\<nu>) \<star> cod (snd \<mu>\<nu>))\<guillemotright>"
proof -
have 1: "VV.arr \<mu>\<nu>"
- using assms VV.arr_char by simp
+ using assms VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show 2: "\<guillemotleft>cmp\<^sub>D\<^sub>N \<mu>\<nu> : DN (dom (fst \<mu>\<nu>)) \<star>\<^sub>B DN (dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B DN (cod (fst \<mu>\<nu>) \<star> cod (snd \<mu>\<nu>))\<guillemotright>"
proof -
have "HoDN_DN.map (VV.dom \<mu>\<nu>) = DN (dom (fst \<mu>\<nu>)) \<star>\<^sub>B DN (dom (snd \<mu>\<nu>))"
using assms 1 DN.FF_def VV.dom_simp by auto
moreover have "DNoH.map (VV.cod \<mu>\<nu>) = DN (cod (fst \<mu>\<nu>) \<star> cod (snd \<mu>\<nu>))"
using assms 1 VV.cod_simp by simp
ultimately show ?thesis
using assms 1 \<Psi>.preserves_hom by auto
qed
show "\<guillemotleft>cmp\<^sub>D\<^sub>N \<mu>\<nu> : DN (src (snd \<mu>\<nu>)) \<rightarrow>\<^sub>B DN (trg (fst \<mu>\<nu>))\<guillemotright>"
using assms 2 B.src_dom [of "cmp\<^sub>D\<^sub>N \<mu>\<nu>"] B.trg_dom [of "cmp\<^sub>D\<^sub>N \<mu>\<nu>"]
by (elim B.in_homE, intro B.in_hhomI) auto
qed
lemma cmp\<^sub>D\<^sub>N_simps [simp]:
assumes "arr (fst \<mu>\<nu>)" and "arr (snd \<mu>\<nu>)" and "src (fst \<mu>\<nu>) = trg (snd \<mu>\<nu>)"
shows "B.arr (cmp\<^sub>D\<^sub>N \<mu>\<nu>)"
and "src\<^sub>B (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (src (snd \<mu>\<nu>))" and "trg\<^sub>B (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (trg (fst \<mu>\<nu>))"
and "B.dom (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (dom (fst \<mu>\<nu>)) \<star>\<^sub>B DN (dom (snd \<mu>\<nu>))"
and "B.cod (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (cod (fst \<mu>\<nu>) \<star> cod (snd \<mu>\<nu>))"
proof
show "VV.arr \<mu>\<nu>"
using assms by blast
have 1: "\<guillemotleft>cmp\<^sub>D\<^sub>N \<mu>\<nu> : DN (src (snd \<mu>\<nu>)) \<rightarrow>\<^sub>B DN (trg (fst \<mu>\<nu>))\<guillemotright>"
using assms by blast
show "src\<^sub>B (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (src (snd \<mu>\<nu>))"
using 1 by fast
show "trg\<^sub>B (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (trg (fst \<mu>\<nu>))"
using 1 by fast
have 2: "\<guillemotleft>cmp\<^sub>D\<^sub>N \<mu>\<nu> : DN (dom (fst \<mu>\<nu>)) \<star>\<^sub>B DN (dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B DN (cod (fst \<mu>\<nu>) \<star> cod (snd \<mu>\<nu>))\<guillemotright>"
using assms by blast
show "B.dom (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (dom (fst \<mu>\<nu>)) \<star>\<^sub>B DN (dom (snd \<mu>\<nu>))"
using 2 by fast
show "B.cod (cmp\<^sub>D\<^sub>N \<mu>\<nu>) = DN (cod (fst \<mu>\<nu>) \<star> cod (snd \<mu>\<nu>))"
using 2 by fast
qed
interpretation DN: pseudofunctor vcomp hcomp \<a> \<i> src trg V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
DN cmp\<^sub>D\<^sub>N
proof
show "\<And>f g h. \<lbrakk> ide f; ide g; ide h; src f = trg g; src g = trg h \<rbrakk> \<Longrightarrow>
DN (\<a> f g h) \<cdot>\<^sub>B cmp\<^sub>D\<^sub>N (f \<star> g, h) \<cdot>\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \<star>\<^sub>B DN h) =
cmp\<^sub>D\<^sub>N (f, g \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \<cdot>\<^sub>B \<a>\<^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 (\<a> f g h) \<cdot>\<^sub>B cmp\<^sub>D\<^sub>N (f \<star> g, h) \<cdot>\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \<star>\<^sub>B DN h) =
cmp\<^sub>D\<^sub>N (f, g \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[DN f, DN g, DN h]"
proof -
have 1: "E.Trg (Dom g) = E.Trg (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<and>
\<lbrace>E.Trg (Dom g)\<rbrace> = \<lbrace>E.Trg (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)\<rbrace>"
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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)
(B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod f \<^bold>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) \<cdot>\<^sub>B
(Map f \<star>\<^sub>B B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)))"
proof -
have "\<guillemotleft>B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod f \<^bold>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) \<cdot>\<^sub>B
(Map f \<star>\<^sub>B
B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
EVAL (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)
\<Rightarrow>\<^sub>B EVAL (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)\<guillemotright>"
proof (intro B.comp_in_homI)
show 2: "\<guillemotleft>B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
EVAL (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<Rightarrow>\<^sub>B
EVAL (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)\<guillemotright>"
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 "\<guillemotleft>B.can (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod f \<^bold>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) :
EVAL (Cod f \<^bold>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) \<Rightarrow>\<^sub>B
EVAL (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)\<guillemotright>"
proof -
have "E.Ide (Cod f \<^bold>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) =
E.Nmlize (Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h" "Cod f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h"]
by blast
qed
show "\<guillemotleft>Map f \<star>\<^sub>B B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
EVAL (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<Rightarrow>\<^sub>B EVAL (Cod f \<^bold>\<star> Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)\<guillemotright>"
using f g h fg gh B.can_in_hom
apply simp
proof (intro B.hcomp_in_vhom B.comp_in_homI)
show 1: "\<guillemotleft>B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
EVAL (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<Rightarrow>\<^sub>B EVAL (Dom g \<^bold>\<star> Dom h)\<guillemotright>"
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 "\<guillemotleft>B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) :
EVAL (Cod g \<^bold>\<star> Cod h) \<Rightarrow>\<^sub>B EVAL (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)\<guillemotright>"
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 "\<guillemotleft>Map g \<star>\<^sub>B Map h : EVAL (Dom g \<^bold>\<star> Dom h) \<Rightarrow>\<^sub>B EVAL (Cod g \<^bold>\<star> Cod h)\<guillemotright>"
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 \<star>\<^sub>B Map h) = EVAL (Dom g \<^bold>\<star> Dom h)"
using g h gh 2 by fastforce
show "B.cod (Map g \<star>\<^sub>B Map h) = EVAL (Cod g \<^bold>\<star> Cod h)"
using g h gh 2 by fastforce
qed
show "\<guillemotleft>Map f : Map f \<Rightarrow>\<^sub>B EVAL (Cod f)\<guillemotright>"
using f arr_char Cod_ide by auto
show "src\<^sub>B (Map f) = trg\<^sub>B \<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>"
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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) = E.Nmlize (Dom g \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) =
E.Nmlize (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 (\<a> f g h) \<cdot>\<^sub>B cmp\<^sub>D\<^sub>N (f \<star> g, h) \<cdot>\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \<star>\<^sub>B DN h) =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> 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
+ using f g h fg gh VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char ideD by simp
have 10: "VV.ide (f, g)"
- using f g fg VV.ide_char by auto
+ using f g fg VV.ide_char\<^sub>S\<^sub>b\<^sub>C by auto
have 11: "VV.ide (hcomp f g, h)"
- using f g h fg gh VV.ide_char VV.arr_char by simp
+ using f g h fg gh VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have 12: "arr (MkArr (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)
(B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B
B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)))"
proof (intro arr_MkArr)
show "Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h \<in> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h \<in> IDE"
using g h gh Cod_ide \<open>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h \<in> IDE\<close> by auto
show "B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B
B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)
\<in> HOM (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)"
proof
show "E.Src (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) = E.Src (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) \<and>
E.Trg (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) = E.Trg (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) \<and>
\<guillemotleft>B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
\<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h\<rbrace>\<guillemotright>"
proof (intro conjI)
show "E.Src (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) = E.Src (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)"
using g h gh Cod_ide by simp
show "E.Trg (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) = E.Trg (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h)"
using g h gh Cod_ide by simp
show "\<guillemotleft>B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
\<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h\<rbrace>\<guillemotright>"
proof (intro B.comp_in_homI)
show "\<guillemotleft>B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
\<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Dom g \<^bold>\<star> Dom h\<rbrace>\<guillemotright>"
using 3 4 5 by blast
show "\<guillemotleft>Map g \<star>\<^sub>B Map h : \<lbrace>Dom g \<^bold>\<star> Dom h\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod g \<^bold>\<star> Cod h\<rbrace>\<guillemotright>"
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 "\<guillemotleft>B.can (Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h) (Cod g \<^bold>\<star> Cod h) :
\<lbrace>Cod g \<^bold>\<star> Cod h\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Cod h\<rbrace>\<guillemotright>"
using 3 4 5 Cod_ide g h by auto
qed
qed
qed
qed
have "DN (\<a> f g h) = \<lbrace>Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>"
proof -
have "DN (\<a> f g h) =
(B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
((Map f \<star>\<^sub>B B.can (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom g \<^bold>\<star> Dom h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h))) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h))"
using f g h fg gh 1 2 9 12 DN_def \<a>_def Cod_ide by simp
also have
"... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
(Map f \<star>\<^sub>B \<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
proof -
have "\<guillemotleft>B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
\<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace> \<Rightarrow>\<^sub>B Map g \<star>\<^sub>B Map h\<guillemotright>"
using g h 3 4 5 B.can_in_hom [of "Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h" "Dom g \<^bold>\<star> Dom h"]
by simp
hence "Map f \<star>\<^sub>B B.can (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom g \<^bold>\<star> Dom h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) =
Map f \<star>\<^sub>B B.can (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom g \<^bold>\<star> Dom h) \<cdot>\<^sub>B
B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
using B.comp_cod_arr by auto
also have "... = Map f \<star>\<^sub>B \<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>"
using f g h fg gh 3 4 5 B.can_Ide_self by auto
finally have "Map f \<star>\<^sub>B B.can (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom g \<^bold>\<star> Dom h) \<cdot>\<^sub>B
(Map g \<star>\<^sub>B Map h) \<cdot>\<^sub>B B.can (Dom g \<^bold>\<star> Dom h) (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) =
Map f \<star>\<^sub>B \<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>"
by simp
thus ?thesis by simp
qed
also have
"... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
proof -
have "\<guillemotleft>B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) :
\<lbrace>Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace> \<Rightarrow>\<^sub>B Map f \<star>\<^sub>B \<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>\<guillemotright>"
using f g h 6 7 8
B.can_in_hom [of "Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h" "Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h"]
by simp
hence "(Map f \<star>\<^sub>B \<lbrace>Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) =
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
using B.comp_cod_arr by auto
thus ?thesis by simp
qed
also have
"... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
using f g h fg gh 6 7 8 by auto
also have "... = \<lbrace>Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>"
using f g h fg gh 6 B.can_Ide_self by blast
finally show ?thesis by simp
qed
have "DN (\<a> f g h) \<cdot>\<^sub>B cmp\<^sub>D\<^sub>N (f \<star> g, h) \<cdot>\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \<star>\<^sub>B DN h) =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
B.can ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) \<cdot>\<^sub>B
(B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) (Dom f \<^bold>\<star> Dom g) \<star>\<^sub>B Map h)"
proof -
have "DN (\<a> f g h) =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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) \<open>DN (\<a> f g h) = \<lbrace>Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h\<rbrace>\<close> ide_char
by (metis (no_types, lifting) arr_char ideD(1))
thus ?thesis
using f g h fg gh 10 11 DN_def \<Psi>.map_simp_ide by simp
qed
also have
"... = (B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
B.can ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h)) \<cdot>\<^sub>B
(B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) (Dom f \<^bold>\<star> Dom g) \<star>\<^sub>B Map h)"
using B.comp_assoc by simp
also have
"... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) \<cdot>\<^sub>B
B.can ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> Dom h)"
proof -
have "B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) (Dom f \<^bold>\<star> Dom g) \<star>\<^sub>B Map h =
B.can ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> Dom h)"
proof -
have "B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) (Dom f \<^bold>\<star> Dom g) \<star>\<^sub>B Map h =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) (Dom f \<^bold>\<star> Dom g) \<star>\<^sub>B B.can (Dom h) (Dom h)"
using h B.can_Ide_self by fastforce
also have "... = B.can ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> 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>\<star> Dom g" "Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
B.can ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h)"
proof -
have "E.Ide ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
using f g h 1 4 6 by blast
moreover have "E.Nmlize ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) =
E.Nmlize ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) =
E.Nmlize (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> Dom h)"
proof -
have "E.Ide ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> Dom h)"
using 1 4 7 by simp
moreover have "E.Ide ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h)"
using f g h 6 by blast
moreover have "E.Nmlize ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> Dom h) =
E.Nmlize ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h)"
using f g h 1 7 E.Nml_HcompNml(1) by fastforce
moreover have "E.Nmlize ((Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g) \<^bold>\<star> Dom h) =
E.Nmlize (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> 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 \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \<cdot>\<^sub>B
\<a>\<^sub>B[DN f, DN g, DN h]"
proof -
have "cmp\<^sub>D\<^sub>N (f, g \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[DN f, DN g, DN h] =
(cmp\<^sub>D\<^sub>N (f, g \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h))) \<cdot>\<^sub>B \<a>\<^sub>B[DN f, DN g, DN h]"
using B.comp_assoc by simp
also have
"... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<star> Dom h) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<star> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> Dom h)"
proof -
have "cmp\<^sub>D\<^sub>N (f, g \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<star> Dom h)"
proof -
have "cmp\<^sub>D\<^sub>N (f, g \<star> h) \<cdot>\<^sub>B (DN f \<star>\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) =
B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
(Map f \<star>\<^sub>B B.can (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom g \<^bold>\<star> Dom h))"
- using f g h fg gh VV.ide_char VV.arr_char DN_def by simp
+ using f g h fg gh VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C DN_def by simp
also have
"... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
(B.can (Dom f) (Dom f) \<star>\<^sub>B B.can (Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom g \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) \<cdot>\<^sub>B
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<star> Dom h)"
using 1 4 5 7 B.hcomp_can by auto
also have "... = B.can (Dom f \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) (Dom f \<^bold>\<star> Dom g \<^bold>\<star> Dom h)"
using 1 4 5 6 7 8 B.vcomp_can by auto
finally show ?thesis by simp
qed
moreover have "\<a>\<^sub>B[DN f, DN g, DN h] =
B.can (Dom f \<^bold>\<star> Dom g \<^bold>\<star> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> 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>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom g \<^bold>\<lfloor>\<^bold>\<star>\<^bold>\<rfloor> Dom h) ((Dom f \<^bold>\<star> Dom g) \<^bold>\<star> 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 \<a> \<i> src trg V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^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 ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
lemma DN_UP:
assumes "B.arr \<mu>"
shows "DN (UP \<mu>) = \<mu>"
using assms UP_def DN_def arr_UP by auto
interpretation DN: equivalence_pseudofunctor
vcomp hcomp \<a> \<i> src trg V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B DN cmp\<^sub>D\<^sub>N
proof
(* DN is locally (but not globally) full. *)
show "\<And>f f' \<nu>. \<lbrakk> ide f; ide f'; src f = src f'; trg f = trg f'; \<guillemotleft>\<nu> : DN f \<Rightarrow>\<^sub>B DN f'\<guillemotright> \<rbrakk>
\<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> DN \<mu> = \<nu>"
proof -
fix f f' \<nu>
assume f: "ide f" and f': "ide f'"
and eq_src: "src f = src f'" and eq_trg: "trg f = trg f'"
and \<nu>: "\<guillemotleft>\<nu> : DN f \<Rightarrow>\<^sub>B DN f'\<guillemotright>"
show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> DN \<mu> = \<nu>"
proof -
let ?\<mu> = "MkArr (Dom f) (Dom f') \<nu>"
have \<mu>: "\<guillemotleft>?\<mu> : f \<Rightarrow> f'\<guillemotright>"
proof
have "Map f = \<lbrace>Dom f\<rbrace>"
using f by simp
have "Map f' = \<lbrace>Dom f'\<rbrace>"
using f' by simp
have "Dom f' = Cod f'"
using f' Cod_ide by simp
show \<mu>: "arr ?\<mu>"
proof -
have "E.Nml (Dom ?\<mu>) \<and> E.Ide (Dom ?\<mu>)"
proof -
have "E.Nml (Dom f) \<and> E.Ide (Dom f)"
using f ide_char arr_char by blast
thus ?thesis
using f by simp
qed
moreover have "E.Nml (Cod ?\<mu>) \<and> E.Ide (Cod ?\<mu>)"
proof -
have "E.Nml (Dom f') \<and> E.Ide (Dom f')"
using f' ide_char arr_char by blast
thus ?thesis
using f' by simp
qed
moreover have "E.Src (Dom ?\<mu>) = E.Src (Cod ?\<mu>)"
using f f' \<nu> arr_char src_def eq_src ideD(1) by auto
moreover have "E.Trg (Dom ?\<mu>) = E.Trg (Cod ?\<mu>)"
using f f' \<nu> arr_char trg_def eq_trg ideD(1) by auto
moreover have "\<guillemotleft>Map ?\<mu> : \<lbrace>Dom ?\<mu>\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Cod ?\<mu>\<rbrace>\<guillemotright>"
proof -
have "\<guillemotleft>\<nu> : \<lbrace>Dom f\<rbrace> \<Rightarrow>\<^sub>B \<lbrace>Dom f'\<rbrace>\<guillemotright>"
using f f' \<nu> 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' \<nu> ide_char arr_char by blast
qed
show "dom ?\<mu> = f"
using f \<mu> dom_char MkArr_Map MkIde_Dom' by simp
show "cod ?\<mu> = f'"
proof -
have "cod ?\<mu> = MkIde (Dom f')"
using \<mu> 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 ?\<mu> = \<nu>"
using \<mu> DN_def by auto
ultimately show ?thesis by blast
qed
qed
(* DN is biessentially surjective on objects. *)
show "\<And>a'. B.obj a' \<Longrightarrow> \<exists>a. obj a \<and> 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>\<langle>a'\<^bold>\<rangle> \<^bold>\<langle>a'\<^bold>\<rangle> a')"
using a' UP_def [of a'] UP.preserves_reflects_arr [of a'] by auto
moreover have "arr (MkArr \<^bold>\<langle>a'\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a'\<^bold>\<rangle>\<^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 "\<exists>a. obj a \<and> B.equivalent_objects (DN.map\<^sub>0 a) a'"
by blast
qed
(* DN is locally essentially surjective. *)
show "\<And>a b g. \<lbrakk> obj a; obj b; \<guillemotleft>g : DN.map\<^sub>0 a \<rightarrow>\<^sub>B DN.map\<^sub>0 b\<guillemotright>; B.ide g \<rbrakk> \<Longrightarrow>
\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> ide f \<and> B.isomorphic (DN f) g"
proof -
fix a b g
assume a: "obj a" and b: "obj b"
and g: "\<guillemotleft>g : DN.map\<^sub>0 a \<rightarrow>\<^sub>B DN.map\<^sub>0 b\<guillemotright>" 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 "\<guillemotleft>UP g : a \<rightarrow> b\<guillemotright>"
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>\<langle>src\<^sub>B g\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>src\<^sub>B g\<^bold>\<rangle>\<^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>\<langle>src\<^sub>B g\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>src\<^sub>B g\<^bold>\<rangle>\<^sub>0 (src\<^sub>B g) = MkArr \<^bold>\<langle>Map a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>Map a\<^bold>\<rangle>\<^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>\<langle>trg\<^sub>B g\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>trg\<^sub>B g\<^bold>\<rangle>\<^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>\<langle>trg\<^sub>B g\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>trg\<^sub>B g\<^bold>\<rangle>\<^sub>0 (trg\<^sub>B g) = MkArr \<^bold>\<langle>Map b\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>Map b\<^bold>\<rangle>\<^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 "\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> ide f \<and> B.isomorphic (DN f) g"
by blast
qed
qed
theorem DN_is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor vcomp hcomp \<a> \<i> src trg V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
DN cmp\<^sub>D\<^sub>N"
..
text \<open>
The following gives an explicit formula for a component of the unit isomorphism of
the pseudofunctor \<open>UP\<close> 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.
\<close>
interpretation UP: equivalence_pseudofunctor
V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \<a> \<i> 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>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a"
proof -
have " MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a = UP.unit a"
proof (intro UP.unit_eqI)
show "B.obj a"
using assms by simp
have 0: "\<guillemotleft>a : a \<Rightarrow>\<^sub>B a\<guillemotright>"
using assms by auto
have 1: "arr (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a)"
apply (unfold arr_char, intro conjI)
using assms by auto
have 2: "arr (MkArr \<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<langle>a\<^bold>\<rangle> a)"
apply (unfold arr_char, intro conjI)
using assms by auto
have 3: "arr (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 a)"
apply (unfold arr_char, intro conjI)
using assms by auto
show "\<guillemotleft>MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a : UP.map\<^sub>0 a \<Rightarrow> UP a\<guillemotright>"
proof
show "arr (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a)" by fact
show "dom (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> 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>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> 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>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a)"
using assms 1 iso_char by auto
show "MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a \<cdot> \<i> (UP.map\<^sub>0 a) =
(UP \<i>\<^sub>B[a] \<cdot> cmp\<^sub>U\<^sub>P (a, a)) \<cdot> (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a \<star> MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a)"
proof -
have "MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a \<cdot> \<i> (UP.map\<^sub>0 a) = MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a"
unfolding \<i>_def UP.map\<^sub>0_def UP_def
using assms 0 1 2 src_def by auto
also have "... = (UP \<i>\<^sub>B[a] \<cdot> cmp\<^sub>U\<^sub>P (a, a)) \<cdot> (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a \<star> MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a)"
proof -
have "(UP \<i>\<^sub>B[a] \<cdot> cmp\<^sub>U\<^sub>P (a, a)) \<cdot> (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a \<star> MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a) =
(MkArr \<^bold>\<langle>a \<star>\<^sub>B a\<^bold>\<rangle> \<^bold>\<langle>a\<^bold>\<rangle> \<i>\<^sub>B[a] \<cdot> MkArr (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) \<^bold>\<langle>a \<star>\<^sub>B a\<^bold>\<rangle> (a \<star>\<^sub>B a))
\<cdot> (MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a \<star> MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> a)"
using assms UP_def cmp\<^sub>U\<^sub>P_ide_simp by auto
also have "... = (MkArr \<^bold>\<langle>a \<star>\<^sub>B a\<^bold>\<rangle> \<^bold>\<langle>a\<^bold>\<rangle> \<i>\<^sub>B[a] \<cdot> MkArr (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) \<^bold>\<langle>a \<star>\<^sub>B a\<^bold>\<rangle> (a \<star>\<^sub>B a))
\<cdot> MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) (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>\<langle>a\<^bold>\<rangle>\<^sub>0"] B.comp_cod_arr
by auto
also have "... = MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> ((\<i>\<^sub>B[a] \<cdot>\<^sub>B (a \<star>\<^sub>B a)) \<cdot>\<^sub>B B.runit' a)"
proof -
have "MkArr \<^bold>\<langle>a \<star>\<^sub>B a\<^bold>\<rangle> \<^bold>\<langle>a\<^bold>\<rangle> \<i>\<^sub>B[a] \<cdot> MkArr (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) \<^bold>\<langle>a \<star>\<^sub>B a\<^bold>\<rangle> (a \<star>\<^sub>B a) =
MkArr (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) \<^bold>\<langle>a\<^bold>\<rangle> (\<i>\<^sub>B[a] \<cdot>\<^sub>B (a \<star>\<^sub>B a))"
using assms
by (intro comp_MkArr arr_MkArr) auto
moreover have "MkArr (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) \<^bold>\<langle>a\<^bold>\<rangle> (\<i>\<^sub>B[a] \<cdot>\<^sub>B (a \<star>\<^sub>B a))
\<cdot> MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 (\<^bold>\<langle>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>a\<^bold>\<rangle>) (B.runit' a) =
MkArr \<^bold>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> ((\<i>\<^sub>B[a] \<cdot>\<^sub>B (a \<star>\<^sub>B a)) \<cdot>\<^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>\<langle>a\<^bold>\<rangle>\<^sub>0 \<^bold>\<langle>a\<^bold>\<rangle> 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 \<open>
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.
\<close>
locale pseudofunctor_into_strict_bicategory =
pseudofunctor +
D: strict_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
begin
lemma image_of_unitor:
assumes "C.ide g"
shows "F \<l>\<^sub>C[g] = (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
and "F \<r>\<^sub>C[g] = (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
and "F (C.lunit' g) = \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (trg\<^sub>C g) \<star>\<^sub>D F g)"
and "F (C.runit' g) = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))"
proof -
show A: "F \<l>\<^sub>C[g] = (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have 1: "\<guillemotleft>(D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g)) :
F (trg\<^sub>C g \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F g\<guillemotright>"
proof
show "\<guillemotleft>D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g : F (trg\<^sub>C g) \<star>\<^sub>D F g \<Rightarrow>\<^sub>D F g\<guillemotright>"
using assms unit_char by (auto simp add: D.hcomp_obj_arr)
show "\<guillemotleft>D.inv (\<Phi> (trg\<^sub>C g, g)) : F (trg\<^sub>C g \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F (trg\<^sub>C g) \<star>\<^sub>D F g\<guillemotright>"
using assms cmp_in_hom(2) D.inv_is_inverse by simp
qed
have "(D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g)) =
F g \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
using 1 D.comp_cod_arr by auto
also have "... = (F \<l>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (trg\<^sub>C g) \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
using assms lunit_coherence [of g] D.strict_lunit by simp
also have "... = F \<l>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D
((unit (trg\<^sub>C g) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g)) \<cdot>\<^sub>D
D.inv (\<Phi> (trg\<^sub>C g, g))"
using D.comp_assoc by simp
also have "... = F \<l>\<^sub>C[g]"
proof -
have "(unit (trg\<^sub>C g) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) = F (trg\<^sub>C g) \<star>\<^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 "\<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g)) = F (trg\<^sub>C g \<star>\<^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 \<r>\<^sub>C[g] = (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
proof -
have 1: "\<guillemotleft>(F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) :
F (g \<star>\<^sub>C src\<^sub>C g) \<Rightarrow>\<^sub>D F g\<guillemotright>"
proof
show "\<guillemotleft>F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g)) : F g \<star>\<^sub>D F (src\<^sub>C g) \<Rightarrow>\<^sub>D F g\<guillemotright>"
using assms unit_char by (auto simp add: D.hcomp_arr_obj)
show "\<guillemotleft>D.inv (\<Phi> (g, src\<^sub>C g)) : F (g \<star>\<^sub>C src\<^sub>C g) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F (src\<^sub>C g)\<guillemotright>"
using assms cmp_in_hom(2) by simp
qed
have "(F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) =
F g \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
using 1 D.comp_cod_arr by auto
also have "... = (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) \<cdot>\<^sub>D
(F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
using assms runit_coherence [of g] D.strict_runit by simp
also have "... = F \<r>\<^sub>C[g] \<cdot>\<^sub>D (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D ((F g \<star>\<^sub>D unit (src\<^sub>C g)) \<cdot>\<^sub>D
(F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
using D.comp_assoc by simp
also have "... = F \<r>\<^sub>C[g]"
proof -
have "(F g \<star>\<^sub>D unit (src\<^sub>C g)) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) = F g \<star>\<^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 "\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) = F (g \<star>\<^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) = \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (trg\<^sub>C g) \<star>\<^sub>D F g)"
proof -
have "F (C.lunit' g) = D.inv (F \<l>\<^sub>C[g])"
using assms C.iso_lunit preserves_inv by simp
also have "... = D.inv ((D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g)))"
using A by simp
also have "... = \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (trg\<^sub>C g) \<star>\<^sub>D F g)"
proof -
have "D.iso (D.inv (\<Phi> (trg\<^sub>C g, g))) \<and> D.inv (D.inv (\<Phi> (trg\<^sub>C g, g))) = \<Phi> (trg\<^sub>C g, g)"
using assms by simp
moreover have "D.iso (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) \<and>
D.inv (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) = unit (trg\<^sub>C g) \<star>\<^sub>D F g"
using assms unit_char by simp
moreover have "D.seq (D.inv (unit (trg\<^sub>C g)) \<star>\<^sub>D F g) (D.inv (\<Phi> (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) = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))"
proof -
have "F (C.runit' g) = D.inv (F \<r>\<^sub>C[g])"
using assms C.iso_runit preserves_inv by simp
also have "... = D.inv ((F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)))"
using B by simp
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))"
proof -
have "D.iso (D.inv (\<Phi> (g, src\<^sub>C g))) \<and> D.inv (D.inv (\<Phi> (g, src\<^sub>C g))) = \<Phi> (g, src\<^sub>C g)"
using assms by simp
moreover have "D.iso (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) \<and>
D.inv (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) = F g \<star>\<^sub>D unit (src\<^sub>C g)"
using assms unit_char by simp
moreover have "D.seq (F g \<star>\<^sub>D D.inv (unit (src\<^sub>C g))) (D.inv (\<Phi> (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 \<a>\<^sub>C[f, g, h] = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
and "F (C.\<a>' f g h) = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
show 1: "F \<a>\<^sub>C[f, g, h] = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
proof -
have 2: "D.seq (\<Phi> (f, g \<star>\<^sub>C h)) ((F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^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 \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
using assms assoc_coherence [of f g h] by blast
moreover have "D.iso (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^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
+ using assms 2 3 C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_char\<^sub>S\<^sub>b\<^sub>C C.VV.cod_char\<^sub>S\<^sub>b\<^sub>C FF_def D.isos_compose
by auto
ultimately have "F \<a>\<^sub>C[f, g, h] =
(\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h])) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h))"
using D.invert_side_of_triangle(2)
[of "\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
"F \<a>\<^sub>C[f, g, h]" "\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)"]
by presburger
also have "... = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
proof -
have "D.inv (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)) =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
proof -
have "D.seq (\<Phi> (f \<star>\<^sub>C g, h)) (\<Phi> (f, g) \<star>\<^sub>D F h)"
using assms by fastforce
thus ?thesis
using assms D.inv_comp by simp
qed
moreover have "(F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] = (F f \<star>\<^sub>D \<Phi> (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.\<a>' f g h) = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
have "F (C.\<a>' f g h) = D.inv (F \<a>\<^sub>C[f, g, h])"
using assms preserves_inv by simp
also have "... = D.inv (\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h)))"
using 1 by simp
also have "... = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
- using assms C.VV.arr_char FF_def D.hcomp_assoc D.comp_assoc
+ using assms C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<open>
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.
\<close>
locale equivalence_in_strict_bicategory =
strict_bicategory +
equivalence_in_bicategory
begin
lemma triangle_right_implies_left:
shows "(g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) = g \<Longrightarrow> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f"
proof -
text \<open>
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}
\<close>
assume 1: "(g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) = g"
have 2: "(inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) = g"
proof -
have "(inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) = inv ((g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g))"
using antipar inv_comp hcomp_assoc by simp
also have "... = g"
using 1 by simp
finally show ?thesis by blast
qed
have "(\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = (\<epsilon> \<star> f) \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) \<star> f) \<cdot> (f \<star> \<eta>)"
proof -
have "(f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) \<star> f) \<cdot> (f \<star> \<eta>) = f \<star> \<eta>"
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 "... = (\<epsilon> \<star> f) \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) \<star> f) \<cdot> (f \<star> \<eta>) \<star> (inv \<eta> \<cdot> \<eta>)"
proof -
have "inv \<eta> \<cdot> \<eta> = 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 "... = (\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>)) \<star> ((f \<star> inv \<eta>) \<cdot> (f \<star> \<eta>))) \<cdot> (f \<star> \<eta>)"
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 "... = (\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>)) \<star>
((f \<star> inv \<eta>) \<cdot> ((inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>))) \<cdot> (f \<star> \<eta>)"
proof -
have "(inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f) = (f \<star> g) \<star> f"
using whisker_right [of f "inv \<epsilon>" \<epsilon>] 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 "... = (((\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>))) \<cdot> (f \<star> g)) \<star>
(((f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>))) \<cdot>
(f \<star> \<eta>)"
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 "... = (((\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>))) \<star> ((f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f))) \<cdot>
((f \<star> g) \<star> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>))) \<cdot>
(f \<star> \<eta>)"
proof -
have 3: "seq (\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>))) (f \<star> g)"
using 2 antipar by auto
moreover have 4: "seq ((f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) ((\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>))"
using antipar hcomp_assoc by auto
ultimately show ?thesis
using interchange by simp
qed
also have "... = ((\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>))) \<star> ((f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f))) \<cdot>
((f \<star> g) \<star> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>)) \<cdot> (f \<star> \<eta>)"
using comp_assoc by presburger
also have "... = ((\<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot>
((f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>)) \<star> f)) \<cdot>
(f \<star> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) \<star> f) \<cdot> (f \<star> \<eta>)"
proof -
have "(\<epsilon> \<cdot> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>))) \<star> ((f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) =
(\<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot> ((f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>)) \<star> f)"
proof -
have "seq \<epsilon> (f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>))"
using 2 antipar by simp
moreover have "seq ((f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> 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 \<epsilon> "f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>)" "(f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)" f]
by simp
qed
moreover have "((f \<star> g) \<star> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>)) \<cdot> (f \<star> \<eta>) =
(f \<star> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) \<star> f) \<cdot> (f \<star> \<eta>)"
proof -
have "((f \<star> g) \<star> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>)) \<cdot> (f \<star> \<eta>) =
(f \<star> g \<star> \<epsilon> \<star> f) \<cdot> (f \<star> (g \<star> f) \<star> \<eta>) \<cdot> (f \<star> \<eta> \<star> src f)"
using antipar comp_assoc hcomp_assoc whisker_left hcomp_arr_obj by simp
also have "... = (f \<star> g \<star> \<epsilon> \<star> f) \<cdot> (f \<star> ((g \<star> f) \<star> \<eta>) \<cdot> (\<eta> \<cdot> src f))"
using antipar comp_arr_dom whisker_left hcomp_arr_obj by simp
also have "... = (f \<star> g \<star> \<epsilon> \<star> f) \<cdot> (f \<star> \<eta> \<star> \<eta>)"
using antipar comp_arr_dom comp_cod_arr hcomp_arr_obj
interchange [of "g \<star> f" \<eta> \<eta> "src f"]
by simp
also have "... = (f \<star> g \<star> \<epsilon> \<star> f) \<cdot> (f \<star> \<eta> \<star> g \<star> f) \<cdot> (f \<star> src f \<star> \<eta>)"
using antipar comp_arr_dom comp_cod_arr whisker_left
interchange [of \<eta> "src f" "g \<star> f" \<eta>]
by simp
also have "... = ((f \<star> g \<star> \<epsilon> \<star> f) \<cdot> (f \<star> \<eta> \<star> g \<star> f)) \<cdot> (f \<star> \<eta>)"
using antipar comp_assoc by (simp add: hcomp_obj_arr)
also have "... = (f \<star> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) \<star> f) \<cdot> (f \<star> \<eta>)"
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 "... = (\<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot>
((f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) \<star> f) \<cdot>
(f \<star> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) \<star> f)) \<cdot> (f \<star> \<eta>)"
using comp_assoc hcomp_assoc antipar(1) antipar(2) by auto
also have "... = (\<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot>
((f \<star> (inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) \<cdot> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) \<star> f)) \<cdot>
(f \<star> \<eta>)"
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 "... = (\<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>)"
proof -
have "(inv \<eta> \<star> g) \<cdot> (g \<star> inv \<epsilon>) \<cdot> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> 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 \<star> inv \<eta>) \<cdot> ((inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>)"
proof -
have "(\<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>) =
(trg f \<cdot> \<epsilon> \<star> (f \<star> inv \<eta>) \<cdot> (inv \<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>)"
using hcomp_obj_arr comp_cod_arr by simp
also have "... = ((trg f \<star> f \<star> inv \<eta>) \<cdot> (\<epsilon> \<star> inv \<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>)"
using antipar hcomp_arr_obj hcomp_assoc interchange by auto
also have "... = (f \<star> inv \<eta>) \<cdot> ((inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>)"
proof -
have "(inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f) = (trg f \<star> inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> trg f \<star> f)"
using hseqI' by (simp add: hcomp_obj_arr)
also have "... = \<epsilon> \<star> inv \<epsilon> \<star> f"
using antipar comp_arr_dom comp_cod_arr
interchange [of "trg f" \<epsilon> "inv \<epsilon> \<star> f" "trg f \<star> f"]
by force
finally have "(inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f) = \<epsilon> \<star> inv \<epsilon> \<star> f" by simp
moreover have "trg f \<star> f \<star> inv \<eta> = f \<star> inv \<eta>"
using hcomp_obj_arr [of "trg f" "f \<star> inv \<eta>"] by fastforce
ultimately have "((trg f \<star> f \<star> inv \<eta>) \<cdot> (\<epsilon> \<star> inv \<epsilon> \<star> f)) \<cdot> (f \<star> \<eta>) =
((f \<star> inv \<eta>) \<cdot> ((inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f))) \<cdot> (f \<star> \<eta>)"
by simp
thus ?thesis
using comp_assoc by simp
qed
finally show ?thesis by simp
qed
also have "... = f \<star> inv \<eta> \<cdot> \<eta>"
proof -
have "(inv \<epsilon> \<star> f) \<cdot> (\<epsilon> \<star> f) = f \<star> g \<star> 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 \<star> 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 \<open>
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.
\<close>
context equivalence_in_bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> 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 \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
..
interpretation E: equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
\<open>S.UP f\<close> \<open>S.UP g\<close>
\<open>S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit (src f)\<close>
\<open>S.inv (UP.unit (trg f)) \<cdot>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)\<close>
using equivalence_in_bicategory_axioms UP.preserves_equivalence [of f g \<eta> \<epsilon>] by auto
interpretation E: equivalence_in_strict_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
\<open>S.UP f\<close> \<open>S.UP g\<close>
\<open>S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit (src f)\<close>
\<open>S.inv (UP.unit (trg f)) \<cdot>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)\<close>
..
lemma UP_triangle:
shows "S.UP ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) =
S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
and "S.UP (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) =
(S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S UP.unit (src g))) \<cdot>\<^sub>S
(S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
and "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) =
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
and "S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) =
(S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
and "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<Longrightarrow>
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) =
(S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
show T1: "S.UP ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) =
S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
proof -
have "S.UP ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) =
S.UP (g \<star> \<epsilon>) \<cdot>\<^sub>S S.UP \<a>[g, f, g] \<cdot>\<^sub>S S.UP (\<eta> \<star> g)"
using antipar by simp
also have "... =
(S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon>) \<cdot>\<^sub>S
((S.inv (S.cmp\<^sub>U\<^sub>P (g, f \<star> g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \<star> g)) \<cdot>\<^sub>S
(S.UP g \<star>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \<cdot>\<^sub>S
(((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g \<star> f, g)))) \<cdot>\<^sub>S
S.cmp\<^sub>U\<^sub>P (g \<star> f, g)) \<cdot>\<^sub>S (S.UP \<eta> \<star>\<^sub>S S.UP g)) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
proof -
have "S.UP \<a>[g, f, g] =
S.cmp\<^sub>U\<^sub>P (g, f \<star> g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g \<star> f, g))"
using ide_left ide_right antipar UP.image_of_associator [of g f g] by simp
moreover have
"S.UP (g \<star> \<epsilon>) =
S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f \<star> g))"
proof -
have "S.seq (S.cmp\<^sub>U\<^sub>P (g, src g)) (S.UP g \<star>\<^sub>S S.UP \<epsilon>)"
using antipar UP.FF_def UP.cmp_in_hom [of g "src g"]
apply (intro S.seqI) by auto
moreover have
"S.UP (g \<star> \<epsilon>) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \<star> g) = S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon>)"
- using antipar UP.\<Phi>.naturality [of "(g, \<epsilon>)"] UP.FF_def VV.arr_char
+ using antipar UP.\<Phi>.naturality [of "(g, \<epsilon>)"] UP.FF_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C
VV.dom_simp VV.cod_simp
by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (g, f \<star> g))"
using antipar by simp
ultimately show ?thesis
using S.invert_side_of_triangle(2)
[of "S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon>)" "S.UP (g \<star> \<epsilon>)"
"S.cmp\<^sub>U\<^sub>P (g, f \<star> g)"] S.comp_assoc
by presburger
qed
moreover have "S.UP (\<eta> \<star> g) =
(S.cmp\<^sub>U\<^sub>P (g \<star> f, g) \<cdot>\<^sub>S (S.UP \<eta> \<star>\<^sub>S S.UP g)) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
proof -
have "S.UP (\<eta> \<star> g) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (trg g, g) =
S.cmp\<^sub>U\<^sub>P (g \<star> f, g) \<cdot>\<^sub>S (S.UP \<eta> \<star>\<^sub>S S.UP g)"
- using antipar UP.\<Phi>.naturality [of "(\<eta>, g)"] UP.FF_def VV.arr_char
+ using antipar UP.\<Phi>.naturality [of "(\<eta>, g)"] UP.FF_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C
VV.dom_simp VV.cod_simp
by simp
moreover have "S.seq (S.cmp\<^sub>U\<^sub>P (g \<star> f, g)) (S.UP \<eta> \<star>\<^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) \<cdot>\<^sub>S
((S.UP g \<star>\<^sub>S S.UP \<epsilon>) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \<cdot>\<^sub>S
((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S (S.UP \<eta> \<star>\<^sub>S S.UP g)) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
proof -
have "(S.inv (S.cmp\<^sub>U\<^sub>P (g \<star> f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (g \<star> f, g)) \<cdot>\<^sub>S (S.UP \<eta> \<star>\<^sub>S S.UP g) =
(S.UP \<eta> \<star>\<^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 \<star> g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \<star> g)) \<cdot>\<^sub>S
(S.UP g \<star>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))
= (S.UP g \<star>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))"
proof -
have "S.inv (S.cmp\<^sub>U\<^sub>P (g, f \<star> g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \<star> g) = S.UP g \<star>\<^sub>S S.UP (f \<star> 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
+ using antipar VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) \<cdot>\<^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
+ using antipar VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.whisker_left S.whisker_right by auto
finally show ?thesis by simp
qed
show T2: "S.UP (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) =
(S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S UP.unit (src g))) \<cdot>\<^sub>S
(S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))"
using UP.image_of_unitor by simp
show "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) =
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) =
S.UP (\<epsilon> \<star> f) \<cdot>\<^sub>S S.UP \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>S S.UP (f \<star> \<eta>)"
using antipar by simp
also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (f \<star> g, f)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f \<star> g, f) \<cdot>\<^sub>S
(S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (f, g \<star> f)) \<cdot>\<^sub>S
S.cmp\<^sub>U\<^sub>P (f, g \<star> f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "S.UP \<a>\<^sup>-\<^sup>1[f, g, f] =
S.cmp\<^sub>U\<^sub>P (f \<star> g, f) \<cdot>\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (f, g \<star> f))"
using ide_left ide_right antipar UP.image_of_associator by simp
moreover have "S.UP (\<epsilon> \<star> f) =
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f \<star> g, f))"
proof -
have "S.seq (S.cmp\<^sub>U\<^sub>P (trg f, f)) (S.UP \<epsilon> \<star>\<^sub>S S.UP f)"
- using antipar UP.FF_def VV.ide_char VV.arr_char UP.cmp_in_hom [of "trg f" f]
+ using antipar UP.FF_def VV.ide_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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) \<cdot>\<^sub>S (S.UP \<epsilon> \<star>\<^sub>S S.UP f) = S.UP (\<epsilon> \<star> f) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f \<star> g, f)"
- using antipar UP.\<Phi>.naturality [of "(\<epsilon>, f)"] UP.FF_def VV.arr_char
+ using antipar UP.\<Phi>.naturality [of "(\<epsilon>, f)"] UP.FF_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C
VV.dom_simp VV.cod_simp
by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (f \<star> 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) \<cdot>\<^sub>S (S.UP \<epsilon> \<star>\<^sub>S S.UP f)" "S.UP (\<epsilon> \<star> f)"
"S.cmp\<^sub>U\<^sub>P (f \<star> g, f)"]
by metis
qed
moreover have "S.UP (f \<star> \<eta>) =
(S.cmp\<^sub>U\<^sub>P (f, g \<star> f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "S.cmp\<^sub>U\<^sub>P (f, g \<star> f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>) =
S.UP (f \<star> \<eta>) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, src f)"
- using antipar UP.\<Phi>.naturality [of "(f, \<eta>)"] UP.FF_def VV.arr_char
+ using antipar UP.\<Phi>.naturality [of "(f, \<eta>)"] UP.FF_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C
VV.dom_simp VV.cod_simp
by simp
moreover have "S.seq (S.cmp\<^sub>U\<^sub>P (f, g \<star> f)) (S.UP f \<star>\<^sub>S S.UP \<eta>)"
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) \<cdot>\<^sub>S
((S.UP \<epsilon> \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
((S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "(S.inv (S.cmp\<^sub>U\<^sub>P (f \<star> g, f)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f \<star> g, f)) \<cdot>\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) =
(S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f)"
- using antipar S.comp_cod_arr VV.arr_char S.comp_inv_arr' by auto
+ using antipar S.comp_cod_arr VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.comp_inv_arr' by auto
moreover have "(S.inv (S.cmp\<^sub>U\<^sub>P (f, g \<star> f)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g \<star> f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.UP \<eta>)
= (S.UP f \<star>\<^sub>S S.UP \<eta>)"
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) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^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
+ using antipar VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.whisker_left S.whisker_right by auto
finally show ?thesis by simp
qed
show "S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) =
(S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
using UP.image_of_unitor by simp
show "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<Longrightarrow>
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) =
(S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
assume A: "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
have B: "(S.UP g \<star>\<^sub>S S.inv (UP.unit (src g)) \<cdot>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit (trg g) \<star>\<^sub>S S.UP g) = S.UP g"
proof -
show "(S.UP g \<star>\<^sub>S S.inv (UP.unit (src g)) \<cdot>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit (trg g) \<star>\<^sub>S S.UP g) = S.UP g"
proof -
have 2: "S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))
= (S.cmp\<^sub>U\<^sub>P (g, src g) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S UP.unit (src g))) \<cdot>\<^sub>S
(S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g) \<cdot>\<^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 \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)))"
- using antipar VV.arr_char S.hcomp_assoc
+ using antipar VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<star>\<^sub>S UP.unit (src g)) \<cdot>\<^sub>S (S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g) \<cdot>\<^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 \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) =
(S.UP g \<star>\<^sub>S UP.unit (src g)) \<cdot>\<^sub>S (S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g)"
proof -
have "((S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S
S.UP \<eta> \<star>\<^sub>S S.UP g)) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)) =
((S.UP g \<star>\<^sub>S UP.unit (src g)) \<cdot>\<^sub>S (S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g)) \<cdot>\<^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 \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S
S.UP \<eta> \<star>\<^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 \<star>\<^sub>S UP.unit (src g)) \<cdot>\<^sub>S (S.inv (UP.unit (trg g)) \<star>\<^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 \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g)"
- using antipar VV.arr_char S.hcomp_assoc by auto
+ using antipar VV.arr_char\<^sub>S\<^sub>b\<^sub>C S.hcomp_assoc by auto
have 3: "(S.UP g \<star>\<^sub>S S.inv (UP.unit (src g))) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g) =
(S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g)"
proof -
have "S.seq (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g)"
using 6 by simp
moreover have "(S.UP g \<star>\<^sub>S UP.unit (src g)) \<cdot>\<^sub>S (S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g) =
(S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g)"
using 5 by argo
moreover have "S.iso (S.UP g \<star>\<^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 \<star>\<^sub>S UP.unit (src g)) =
S.UP g \<star>\<^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 \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S
(S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g)"
"S.UP g \<star>\<^sub>S UP.unit (src g)" "S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g"]
by presburger
qed
have 4: "((S.UP g \<star>\<^sub>S S.inv (UP.unit (src g))) \<cdot>\<^sub>S
(S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \<cdot>\<^sub>S
((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g)) \<cdot>\<^sub>S
(UP.unit (trg g) \<star>\<^sub>S S.UP g)
= S.UP g"
proof -
have "(((S.UP g \<star>\<^sub>S S.inv (UP.unit (src g))) \<cdot>\<^sub>S (S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \<cdot>\<^sub>S
((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g)) \<cdot>\<^sub>S
(UP.unit (trg g) \<star>\<^sub>S S.UP g)) =
(((S.UP g \<star>\<^sub>S S.inv (UP.unit (src g))) \<cdot>\<^sub>S
(S.UP g \<star>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \<cdot>\<^sub>S
((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<star>\<^sub>S S.UP g))) \<cdot>\<^sub>S
(UP.unit (trg g) \<star>\<^sub>S S.UP g)"
using S.comp_assoc by simp
also have "... =
(S.inv (UP.unit (trg g)) \<star>\<^sub>S S.UP g) \<cdot>\<^sub>S (UP.unit (trg g) \<star>\<^sub>S S.UP g)"
using 3 S.comp_assoc by auto
also have "... = S.inv (UP.unit (trg g)) \<cdot>\<^sub>S UP.unit (trg g) \<star>\<^sub>S S.UP g"
using UP.unit_char(2) S.whisker_right by auto
also have "... = UP.map\<^sub>0 (trg g) \<star>\<^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) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) =
(S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "(S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) =
(UP.unit (trg f) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have 2: "(S.inv (UP.unit (trg f)) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
((S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit (src f)) =
S.UP f"
proof -
have "S.UP f = (S.inv (UP.unit (trg f)) \<cdot>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S
S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit (src f))"
using B antipar E.triangle_right_implies_left by argo
also have "... = (S.inv (UP.unit (trg f)) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
((S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit (src f))"
proof -
have "S.inv (UP.unit (trg f)) \<cdot>\<^sub>S S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f =
(S.inv (UP.unit (trg f)) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f)"
using UP.unit_char S.whisker_right by simp
moreover have "S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit (src f) =
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^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 \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit (src f)) =
(UP.unit (trg f) \<star>\<^sub>S S.UP f)"
proof -
have "S.inv (S.inv (UP.unit (trg f)) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.UP f =
UP.unit (trg f) \<star>\<^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)) \<star>\<^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)) \<star>\<^sub>S S.UP f"
"((S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit (src f))"]
by presburger
qed
moreover have "S.hseq (UP.unit (trg f)) (S.UP f) \<and>
S.iso (S.UP f \<star>\<^sub>S UP.unit (src f)) \<and>
S.inv (S.UP f \<star>\<^sub>S UP.unit (src f)) = S.UP f \<star>\<^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) \<star>\<^sub>S S.UP f"
"(S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)"
"S.UP f \<star>\<^sub>S UP.unit (src f)"]
by presburger
qed
qed
hence "S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S ((S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) =
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S ((UP.unit (trg f) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))) \<cdot>\<^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 \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
shows "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "par ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
using antipar by simp
moreover have "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) = S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
using assms UP_triangle(3-5) by simp
ultimately show "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
using UP.is_faithful by blast
qed
text \<open>
We \emph{really} don't want to go through the ordeal of proving a dual version of
\<open>UP_triangle(5)\<close>, do we? So let's be smart and dualize via the opposite bicategory.
\<close>
lemma triangle_left_implies_right:
assumes "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
shows "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
proof -
interpret Cop: op_bicategory V H \<a> \<i> src trg ..
interpret Eop: equivalence_in_bicategory V Cop.H Cop.\<a> \<i> Cop.src Cop.trg g f \<eta> \<epsilon>
using antipar unit_in_hom counit_in_hom
by (unfold_locales) simp_all
have "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<Longrightarrow>
(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[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
+ VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C Eop.triangle_right_implies_left
by simp
thus ?thesis
using assms by simp
qed
lemma triangle_left_iff_right:
shows "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<longleftrightarrow>
(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using triangle_left_implies_right triangle_right_implies_left by auto
end
text \<open>
We might as well specialize the dual result back to the strict case while we are at it.
\<close>
context equivalence_in_strict_bicategory
begin
lemma triangle_left_iff_right:
shows "(\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f \<longleftrightarrow> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) = g"
proof -
have "(\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f \<longleftrightarrow> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] = f"
using strict_lunit strict_runit by simp
moreover have "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>)"
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 "... \<longleftrightarrow> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using triangle_left_iff_right by blast
also have "... \<longleftrightarrow> (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) = g"
proof -
have "\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] = g"
using strict_lunit strict_runit by simp
moreover have "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> 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,1490 +1,1490 @@
(* Title: Subbicategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Sub-Bicategories"
text \<open>
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.
\<close>
theory Subbicategory
imports Bicategory
begin
subsection "Construction"
locale subbicategory =
B: bicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B +
subcategory V Arr
for V :: "'a comp" (infixr "\<cdot>\<^sub>B" 55)
and H :: "'a comp" (infixr "\<star>\<^sub>B" 55)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
and Arr :: "'a \<Rightarrow> bool" +
assumes src_closed: "Arr f \<Longrightarrow> Arr (src\<^sub>B f)"
and trg_closed: "Arr f \<Longrightarrow> Arr (trg\<^sub>B f)"
and hcomp_closed: "\<lbrakk> Arr f; Arr g; trg\<^sub>B f = src\<^sub>B g \<rbrakk> \<Longrightarrow> Arr (g \<star>\<^sub>B f)"
and assoc_closed: "\<lbrakk> Arr f \<and> B.ide f; Arr g \<and> B.ide g; Arr h \<and> B.ide h;
src\<^sub>B f = trg\<^sub>B g; src\<^sub>B g = trg\<^sub>B h \<rbrakk> \<Longrightarrow> Arr (\<a>\<^sub>B f g h)"
and assoc'_closed: "\<lbrakk> Arr f \<and> B.ide f; Arr g \<and> B.ide g; Arr h \<and> B.ide h;
src\<^sub>B f = trg\<^sub>B g; src\<^sub>B g = trg\<^sub>B h \<rbrakk> \<Longrightarrow> Arr (B.inv (\<a>\<^sub>B f g h))"
and lunit_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.\<ll> f)"
and lunit'_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.inv (B.\<ll> f))"
and runit_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.\<rr> f)"
and runit'_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.inv (B.\<rr> f))"
begin
notation B.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>B _\<guillemotright>")
notation comp (infixr "\<cdot>" 55)
definition hcomp (infixr "\<star>" 53)
where "g \<star> f = (if arr f \<and> arr g \<and> trg\<^sub>B f = src\<^sub>B g then g \<star>\<^sub>B f else null)"
definition src
where "src \<mu> = (if arr \<mu> then src\<^sub>B \<mu> else null)"
definition trg
where "trg \<mu> = (if arr \<mu> then trg\<^sub>B \<mu> else null)"
interpretation src: endofunctor \<open>(\<cdot>)\<close> src
- using src_def null_char inclusion arr_char src_closed trg_closed dom_closed cod_closed
+ using src_def null_char inclusion arr_char\<^sub>S\<^sub>b\<^sub>C 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)
+ by (metis B.src.as_nat_trans.preserves_comp_2 comp_char seq_char\<^sub>S\<^sub>b\<^sub>C)
interpretation trg: endofunctor \<open>(\<cdot>)\<close> trg
- using trg_def null_char inclusion arr_char src_closed trg_closed dom_closed cod_closed
+ using trg_def null_char inclusion arr_char\<^sub>S\<^sub>b\<^sub>C 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)
+ by (metis B.trg.as_nat_trans.preserves_comp_2 comp_char seq_char\<^sub>S\<^sub>b\<^sub>C)
interpretation horizontal_homs \<open>(\<cdot>)\<close> src trg
- using src_def trg_def src.preserves_arr trg.preserves_arr null_char ide_char arr_char
+ using src_def trg_def src.preserves_arr trg.preserves_arr null_char ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
inclusion
by (unfold_locales, simp_all)
interpretation "functor" VV.comp \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
- using hcomp_def VV.arr_char src_def trg_def arr_char hcomp_closed dom_char cod_char
- VV.dom_char VV.cod_char
+ using hcomp_def VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_closed dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C
+ VV.dom_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C
apply unfold_locales
apply auto[2]
proof -
fix f
assume f: "VV.arr f"
show "dom (fst f \<star> snd f) = fst (VV.dom f) \<star> snd (VV.dom f)"
proof -
have "dom (fst f \<star> snd f) = B.dom (fst f) \<star>\<^sub>B B.dom (snd f)"
proof -
have "dom (fst f \<star> snd f) = B.dom (fst f \<star> snd f)"
- using f dom_char
- by (simp add: arr_char hcomp_closed hcomp_def)
+ using f dom_char\<^sub>S\<^sub>b\<^sub>C
+ by (simp add: arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_closed hcomp_def)
also have "... = B.dom (fst f) \<star>\<^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) \<star> snd (VV.dom f)"
- using f VV.arr_char VV.dom_char arr_char hcomp_def B.seq_if_composable dom_closed
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.dom_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def B.seq_if_composable dom_closed
apply simp
- by (metis (no_types, lifting) dom_char)
+ by (metis (no_types, lifting) dom_char\<^sub>S\<^sub>b\<^sub>C)
finally show ?thesis by simp
qed
show "cod (fst f \<star> snd f) = fst (VV.cod f) \<star> snd (VV.cod f)"
proof -
have "cod (fst f \<star> snd f) = B.cod (fst f) \<star>\<^sub>B B.cod (snd f)"
- using f VV.arr_char arr_char cod_char hcomp_def src_def trg_def
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C 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) \<star> snd (VV.cod f)"
- using f VV.arr_char VV.cod_char arr_char hcomp_def B.seq_if_composable cod_closed
+ using f VV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def B.seq_if_composable cod_closed
apply simp
- by (metis (no_types, lifting) cod_char)
+ by (metis (no_types, lifting) cod_char\<^sub>S\<^sub>b\<^sub>C)
finally show ?thesis by simp
qed
next
fix f g
assume fg: "VV.seq g f"
show "fst (VV.comp g f) \<star> snd (VV.comp g f) = (fst g \<star> snd g) \<cdot> (fst f \<star> snd f)"
proof -
have "fst (VV.comp g f) \<star> snd (VV.comp g f) = fst g \<cdot> fst f \<star> snd g \<cdot> snd f"
- using fg VV.seq_char VV.comp_char VxV.comp_char VxV.not_Arr_Null
+ using fg VV.seq_char\<^sub>S\<^sub>b\<^sub>C 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 \<cdot>\<^sub>B fst f) \<star>\<^sub>B (snd g \<cdot>\<^sub>B snd f)"
- using fg comp_char hcomp_def VV.seq_char inclusion arr_char seq_char B.hseq_char
+ using fg comp_char hcomp_def VV.seq_char\<^sub>S\<^sub>b\<^sub>C inclusion arr_char\<^sub>S\<^sub>b\<^sub>C seq_char\<^sub>S\<^sub>b\<^sub>C B.hseq_char
by (metis (no_types, lifting) B.hseq_char' VxV.seq_char null_char)
also have 1: "... = (fst g \<star>\<^sub>B snd g) \<cdot>\<^sub>B (fst f \<star>\<^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)
+ by (metis (no_types, lifting) VV.arrE VV.seq_char\<^sub>S\<^sub>b\<^sub>C fg src_def trg_def)
thus ?thesis
- using fg VV.seq_char VV.arr_char arr_char seq_char inclusion B.interchange
+ using fg VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C seq_char\<^sub>S\<^sub>b\<^sub>C inclusion B.interchange
by (meson VxV.seqE)
qed
also have "... = (fst g \<star> snd g) \<cdot> (fst f \<star> snd f)"
- using fg comp_char hcomp_def VV.seq_char VV.arr_char arr_char seq_char inclusion
+ using fg comp_char hcomp_def VV.seq_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C seq_char\<^sub>S\<^sub>b\<^sub>C 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 \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> src trg
- using arr_char src_def trg_def src_closed trg_closed
+ using arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def src_closed trg_closed
apply (unfold_locales)
using hcomp_def inclusion not_arr_null by auto
abbreviation \<a>
where "\<a> \<mu> \<nu> \<tau> \<equiv> if VVV.arr (\<mu>, \<nu>, \<tau>) then \<a>\<^sub>B \<mu> \<nu> \<tau> else null"
abbreviation (input) \<alpha>\<^sub>S\<^sub>B
where "\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> \<equiv> \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))"
lemma assoc_closed':
assumes "VVV.arr \<mu>\<nu>\<tau>"
shows "Arr (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>)"
proof -
have 1: "B.VVV.arr \<mu>\<nu>\<tau>"
- using assms VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char arr_char
+ using assms VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
src_def trg_def inclusion
by auto
show "Arr (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>)"
proof -
have "Arr (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) =
Arr ((fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>)) \<cdot>\<^sub>B \<alpha>\<^sub>S\<^sub>B (B.VVV.dom \<mu>\<nu>\<tau>))"
using assms 1 B.\<alpha>_def B.assoc_is_natural_1 [of "fst \<mu>\<nu>\<tau>" "fst (snd \<mu>\<nu>\<tau>)" "snd (snd \<mu>\<nu>\<tau>)"]
- VV.arr_char VVV.arr_char B.VVV.arr_char B.VV.arr_char B.VVV.dom_char B.VV.dom_char
+ VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.dom_char\<^sub>S\<^sub>b\<^sub>C
apply simp
- by (metis (no_types, lifting) arr_char dom_char dom_closed src.preserves_dom
+ by (metis (no_types, lifting) arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C dom_closed src.preserves_dom
trg.preserves_dom)
also have "..."
proof (intro comp_closed)
show "Arr (fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>))"
- using assms 1 B.VVV.arr_char B.VV.arr_char hcomp_closed
+ using assms 1 B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_closed
by (metis (no_types, lifting) B.H.preserves_reflects_arr B.trg_hcomp
- VV.arr_char VVV.arrE arr_char)
+ VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arrE arr_char\<^sub>S\<^sub>b\<^sub>C)
show "B.cod (\<a> (fst (B.VVV.dom \<mu>\<nu>\<tau>)) (fst (snd (B.VVV.dom \<mu>\<nu>\<tau>)))
(snd (snd (B.VVV.dom \<mu>\<nu>\<tau>)))) =
B.dom (fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>))"
- using assms 1 VVV.arr_char VV.arr_char B.VxVxV.dom_char
+ using assms 1 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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.\<alpha>.preserves_reflects_arr
- B.assoc_is_natural_1 B.seqE arr_dom dom_char src_dom trg_dom)
+ by (metis (no_types, lifting) B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arrE B.\<alpha>.preserves_reflects_arr
+ B.assoc_is_natural_1 B.seqE arr_dom dom_char\<^sub>S\<^sub>b\<^sub>C src_dom trg_dom)
show "Arr (\<a> (fst (B.VVV.dom \<mu>\<nu>\<tau>)) (fst (snd (B.VVV.dom \<mu>\<nu>\<tau>)))
(snd (snd (B.VVV.dom \<mu>\<nu>\<tau>))))"
proof -
have "VVV.arr (B.VVV.dom \<mu>\<nu>\<tau>)"
- using 1 B.VVV.dom_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char
+ using 1 B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply simp
by (metis (no_types, lifting) VVV.arrE arr_dom assms dom_simp src_dom trg_dom)
moreover have "Arr (\<a>\<^sub>B (B.dom (fst \<mu>\<nu>\<tau>)) (B.dom (fst (snd \<mu>\<nu>\<tau>)))
(B.dom (snd (snd \<mu>\<nu>\<tau>))))"
proof -
have "B.VVV.ide (B.VVV.dom \<mu>\<nu>\<tau>)"
using 1 B.VVV.ide_dom by blast
thus ?thesis
- using assms B.\<alpha>_def B.VVV.arr_char B.VV.arr_char B.VVV.ide_char B.VV.ide_char
+ using assms B.\<alpha>_def B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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)
+ by (metis (no_types, lifting) "1" B.ide_dom B.src_dom B.trg_dom VV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ VVV.arrE arr_char\<^sub>S\<^sub>b\<^sub>C)
qed
ultimately show ?thesis
- using 1 B.VVV.ide_dom assoc_closed B.VVV.dom_char
+ using 1 B.VVV.ide_dom assoc_closed B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C
apply simp
- by (metis (no_types, lifting) B.VV.arr_char B.VVV.arrE B.VVV.inclusion
+ by (metis (no_types, lifting) B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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.\<ll> f)"
proof -
have 1: "arr f \<and> arr (B.\<ll> (B.dom f))"
- using assms arr_char lunit_closed dom_closed B.ide_dom inclusion by simp
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C lunit_closed dom_closed B.ide_dom inclusion by simp
moreover have "B.dom f = B.cod (B.\<ll> (B.dom f))"
- using 1 arr_char B.\<ll>.preserves_cod inclusion by simp
+ using 1 arr_char\<^sub>S\<^sub>b\<^sub>C B.\<ll>.preserves_cod inclusion by simp
moreover have "B.\<ll> f = f \<cdot> B.\<ll> (B.dom f)"
- using assms 1 B.\<ll>.is_natural_1 inclusion comp_char arr_char by simp
+ using assms 1 B.\<ll>.is_natural_1 inclusion comp_char arr_char\<^sub>S\<^sub>b\<^sub>C by simp
ultimately show ?thesis
- using arr_char comp_closed cod_char seqI dom_simp by auto
+ using arr_char\<^sub>S\<^sub>b\<^sub>C comp_closed cod_char\<^sub>S\<^sub>b\<^sub>C seqI dom_simp by auto
qed
lemma runit_closed':
assumes "Arr f"
shows "Arr (B.\<rr> f)"
proof -
have 1: "arr f \<and> arr (B.\<rr> (B.dom f))"
- using assms arr_char runit_closed dom_closed B.ide_dom inclusion
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C runit_closed dom_closed B.ide_dom inclusion
by simp
moreover have "B.dom f = B.cod (B.\<rr> (B.dom f))"
- using 1 arr_char B.\<ll>.preserves_cod inclusion by simp
+ using 1 arr_char\<^sub>S\<^sub>b\<^sub>C B.\<ll>.preserves_cod inclusion by simp
moreover have "B.\<rr> f = f \<cdot> B.\<rr> (B.dom f)"
- using assms 1 B.\<rr>.is_natural_1 inclusion comp_char arr_char by simp
+ using assms 1 B.\<rr>.is_natural_1 inclusion comp_char arr_char\<^sub>S\<^sub>b\<^sub>C by simp
ultimately show ?thesis
- using arr_char comp_closed cod_char seqI dom_simp by auto
+ using arr_char\<^sub>S\<^sub>b\<^sub>C comp_closed cod_char\<^sub>S\<^sub>b\<^sub>C seqI dom_simp by auto
qed
interpretation natural_isomorphism VVV.comp \<open>(\<cdot>)\<close> HoHV HoVH \<alpha>\<^sub>S\<^sub>B
proof
fix \<mu>\<nu>\<tau>
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> \<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> = null"
by simp
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have 1: "B.VVV.arr \<mu>\<nu>\<tau>"
- using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char arr_char
+ using \<mu>\<nu>\<tau> VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
src_def trg_def inclusion
by auto
show "dom (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = HoHV (VVV.dom \<mu>\<nu>\<tau>)"
proof -
have "dom (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = B.HoHV (B.VVV.dom \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> 1 arr_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
+ using \<mu>\<nu>\<tau> 1 arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
B.\<alpha>_def assoc_closed' dom_simp
by simp
also have "... = HoHV (VVV.dom \<mu>\<nu>\<tau>)"
proof -
have "HoHV (VVV.dom \<mu>\<nu>\<tau>) = HoHV (VxVxV.dom \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> VVV.dom_char VV.arr_char src_def trg_def VVV.arr_char by auto
+ using \<mu>\<nu>\<tau> VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
also have "... = B.HoHV (B.VVV.dom \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> 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
+ using \<mu>\<nu>\<tau> VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
+ HoHV_def B.HoHV_def arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.dom_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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 (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = HoVH (VVV.cod \<mu>\<nu>\<tau>)"
proof -
have "cod (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = B.HoVH (B.VVV.cod \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> 1 arr_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
+ using \<mu>\<nu>\<tau> 1 arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
B.\<alpha>_def assoc_closed' cod_simp
by simp
also have "... = HoVH (VVV.cod \<mu>\<nu>\<tau>)"
proof -
have "HoVH (VVV.cod \<mu>\<nu>\<tau>) = HoVH (VxVxV.cod \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> VVV.cod_char VV.arr_char src_def trg_def VVV.arr_char by auto
+ using \<mu>\<nu>\<tau> VVV.cod_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C by auto
also have "... = B.HoVH (B.VVV.cod \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> 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
+ using \<mu>\<nu>\<tau> VVV.cod_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ HoVH_def B.HoVH_def arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.cod_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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 \<mu>\<nu>\<tau>) \<and> Arr (fst (snd \<mu>\<nu>\<tau>)) \<and> Arr (snd (snd \<mu>\<nu>\<tau>)) \<and>
src\<^sub>B (fst \<mu>\<nu>\<tau>) = trg\<^sub>B (fst (snd \<mu>\<nu>\<tau>)) \<and>
src\<^sub>B (fst (snd \<mu>\<nu>\<tau>)) = trg\<^sub>B (snd (snd \<mu>\<nu>\<tau>))"
- using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char src_def trg_def arr_char by auto
+ using \<mu>\<nu>\<tau> VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C by auto
show "HoVH \<mu>\<nu>\<tau> \<cdot> \<alpha>\<^sub>S\<^sub>B (VVV.dom \<mu>\<nu>\<tau>) = \<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>"
proof -
have "\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> = (fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>)) \<cdot>\<^sub>B
\<a>\<^sub>B (B.dom (fst \<mu>\<nu>\<tau>)) (B.dom (fst (snd \<mu>\<nu>\<tau>))) (B.dom (snd (snd \<mu>\<nu>\<tau>)))"
using 3 inclusion B.assoc_is_natural_1 [of "fst \<mu>\<nu>\<tau>" "fst (snd \<mu>\<nu>\<tau>)" "snd (snd \<mu>\<nu>\<tau>)"]
by (simp add: \<mu>\<nu>\<tau>)
also have "... = (fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>) \<star> snd (snd \<mu>\<nu>\<tau>)) \<cdot>
\<a>\<^sub>B (dom (fst \<mu>\<nu>\<tau>)) (dom (fst (snd \<mu>\<nu>\<tau>))) (dom (snd (snd \<mu>\<nu>\<tau>)))"
using 1 3 \<mu>\<nu>\<tau> hcomp_closed assoc_closed dom_closed hcomp_def comp_def inclusion
- comp_char dom_char VVV.arr_char VV.arr_char
+ comp_char dom_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply simp
- using B.hcomp_simps(2-3) arr_char by presburger
+ using B.hcomp_simps(2-3) arr_char\<^sub>S\<^sub>b\<^sub>C by presburger
also have "... = HoVH \<mu>\<nu>\<tau> \<cdot> \<alpha>\<^sub>S\<^sub>B (VVV.dom \<mu>\<nu>\<tau>)"
- using \<mu>\<nu>\<tau> B.\<alpha>_def HoVH_def VVV.dom_char VV.dom_char VxVxV.dom_char
+ using \<mu>\<nu>\<tau> B.\<alpha>_def HoVH_def VVV.dom_char\<^sub>S\<^sub>b\<^sub>C VV.dom_char\<^sub>S\<^sub>b\<^sub>C VxVxV.dom_char
apply simp
- by (metis (no_types, lifting) VV.arr_char VVV.arrE VVV.arr_dom VxV.dom_char
+ by (metis (no_types, lifting) VV.arr_char\<^sub>S\<^sub>b\<^sub>C VVV.arrE VVV.arr_dom VxV.dom_char
dom_simp)
finally show ?thesis by argo
qed
show "\<alpha>\<^sub>S\<^sub>B (VVV.cod \<mu>\<nu>\<tau>) \<cdot> HoHV \<mu>\<nu>\<tau> = \<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>"
proof -
have "\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> =
\<a>\<^sub>B (B.cod (fst \<mu>\<nu>\<tau>)) (B.cod (fst (snd \<mu>\<nu>\<tau>))) (B.cod (snd (snd \<mu>\<nu>\<tau>))) \<cdot>\<^sub>B
(fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>)) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>)"
using 3 inclusion B.assoc_is_natural_2 [of "fst \<mu>\<nu>\<tau>" "fst (snd \<mu>\<nu>\<tau>)" "snd (snd \<mu>\<nu>\<tau>)"]
by (simp add: \<mu>\<nu>\<tau>)
also have "... = \<a>\<^sub>B (cod (fst \<mu>\<nu>\<tau>)) (cod (fst (snd \<mu>\<nu>\<tau>))) (cod (snd (snd \<mu>\<nu>\<tau>))) \<cdot>
((fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>)) \<star> snd (snd \<mu>\<nu>\<tau>))"
- by (simp add: "3" arrI assoc_closed cod_char cod_closed hcomp_closed hcomp_def
+ by (simp add: "3" arrI\<^sub>S\<^sub>b\<^sub>C assoc_closed cod_char\<^sub>S\<^sub>b\<^sub>C cod_closed hcomp_closed hcomp_def
inclusion comp_def)
also have "... = \<alpha>\<^sub>S\<^sub>B (VVV.cod \<mu>\<nu>\<tau>) \<cdot> HoHV \<mu>\<nu>\<tau>"
- using \<mu>\<nu>\<tau> B.\<alpha>_def HoHV_def VVV.cod_char VV.cod_char VxVxV.cod_char
- VVV.arr_char VV.arr_char arr_cod src_cod trg_cod
+ using \<mu>\<nu>\<tau> B.\<alpha>_def HoHV_def VVV.cod_char\<^sub>S\<^sub>b\<^sub>C VV.cod_char\<^sub>S\<^sub>b\<^sub>C VxVxV.cod_char
+ VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_cod src_cod trg_cod
by simp
finally show ?thesis by argo
qed
next
fix fgh
assume fgh: "VVV.ide fgh"
show "iso (\<alpha>\<^sub>S\<^sub>B fgh)"
proof -
have 1: "B.arr (fst (snd fgh)) \<and> B.arr (snd (snd fgh)) \<and>
src\<^sub>B (fst (snd fgh)) = trg\<^sub>B (snd (snd fgh)) \<and>
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
+ using fgh VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
+ arr_char\<^sub>S\<^sub>b\<^sub>C inclusion
by auto
have 2: "B.ide (fst fgh) \<and> B.ide (fst (snd fgh)) \<and> B.ide (snd (snd fgh))"
- using fgh VVV.ide_char ide_char by blast
+ using fgh VVV.ide_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C by blast
have "\<alpha>\<^sub>S\<^sub>B fgh = \<a>\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))"
using fgh B.\<alpha>_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
+ using fgh 1 2 VVV.ide_char B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C
+ src_def trg_def inclusion arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
moreover have "Arr (\<a>\<^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.\<alpha>_def
+ using fgh 1 VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
+ arr_char\<^sub>S\<^sub>b\<^sub>C assoc_closed' B.\<alpha>_def
by simp
moreover have "Arr (B.inv (\<a>\<^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))
+ using fgh 1 VVV.ide_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
+ arr_char\<^sub>S\<^sub>b\<^sub>C assoc'_closed
+ by (simp add: VVV.arr_char\<^sub>S\<^sub>b\<^sub>C "2" B.VVV.ide_char calculation(2))
ultimately show ?thesis
- using fgh iso_char B.\<alpha>.components_are_iso by auto
+ using fgh iso_char\<^sub>S\<^sub>b\<^sub>C B.\<alpha>.components_are_iso by auto
qed
qed
interpretation L: endofunctor \<open>(\<cdot>)\<close> L
using endofunctor_L by auto
interpretation R: endofunctor \<open>(\<cdot>)\<close> R
using endofunctor_R by auto
interpretation L: faithful_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> 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
+ using par inclusion arr_char\<^sub>S\<^sub>b\<^sub>C dom_simp cod_simp by fastforce
moreover have "B.L f = B.L f'"
proof -
have "\<forall>a. Arr a \<longrightarrow> B.arr a"
by (simp add: inclusion)
moreover have 1: "\<forall>a. arr a \<longrightarrow> (if arr a then hseq (trg a) a else arr null)"
using L.preserves_arr by presburger
moreover have "Arr f \<and> Arr (trg f) \<and> trg\<^sub>B f = src\<^sub>B (trg f)"
by (simp add: \<open>B.par f f'\<close> arrE par trg_closed trg_def)
ultimately show ?thesis
by (metis \<open>B.par f f'\<close> 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 \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> L
proof
fix f f' \<nu>
assume f: "ide f" and f': "ide f'" and \<nu>: "\<guillemotleft>\<nu> : L f \<Rightarrow> L f'\<guillemotright>"
have 1: "L f = trg\<^sub>B f \<star>\<^sub>B f \<and> L f' = trg\<^sub>B f' \<star>\<^sub>B f'"
- using f f' hcomp_def trg_def arr_char ide_char trg_closed by simp
+ using f f' hcomp_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C trg_closed by simp
have 2: "\<guillemotleft>\<nu> : trg\<^sub>B f \<star>\<^sub>B f \<Rightarrow>\<^sub>B trg\<^sub>B f' \<star>\<^sub>B f'\<guillemotright>"
using 1 f f' \<nu> 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)
+ dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C hseq_char' arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C trg_closed null_char
+ by (simp add: arr_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C)
show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> L \<mu> = \<nu>"
proof -
let ?\<mu> = "B.\<ll> f' \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B B.inv (B.\<ll> f)"
have \<mu>: "\<guillemotleft>?\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
proof -
have "\<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
- using f f' \<nu> 2 B.\<ll>_ide_simp lunit'_closed lunit_closed' ide_char by auto
+ using f f' \<nu> 2 B.\<ll>_ide_simp lunit'_closed lunit_closed' ide_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis
- using f f' \<nu> in_hom_char arr_char comp_closed ide_char
+ using f f' \<nu> in_hom_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C comp_closed ide_char\<^sub>S\<^sub>b\<^sub>C
lunit'_closed lunit_closed
by (metis (no_types, lifting) B.arrI B.seqE in_homE)
qed
have \<mu>_eq: "?\<mu> = B.\<ll> f' \<cdot> \<nu> \<cdot> B.inv (B.\<ll> f)"
proof -
have "?\<mu> = B.\<ll> f' \<cdot> \<nu> \<cdot>\<^sub>B B.inv (B.\<ll> f)"
by (metis (no_types, lifting) B.arrI B.seqE \<mu> \<nu> arrE comp_closed f f'
- ide_char in_hom_char local.comp_def lunit'_closed lunit_closed)
+ ide_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C comp_def lunit'_closed lunit_closed)
also have "... = B.\<ll> f' \<cdot> \<nu> \<cdot> B.inv (B.\<ll> f)"
- using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
+ using f f' \<nu> \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C inclusion comp_char comp_closed ide_char\<^sub>S\<^sub>b\<^sub>C
lunit'_closed lunit_closed
by (metis (no_types, lifting) B.seqE in_homE)
finally show ?thesis by simp
qed
have "L ?\<mu> = \<nu>"
proof -
have "L ?\<mu> = trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>"
- using \<mu> \<mu>_eq hcomp_def trg_def inclusion arr_char trg_closed by auto
+ using \<mu> \<mu>_eq hcomp_def trg_def inclusion arr_char\<^sub>S\<^sub>b\<^sub>C trg_closed by auto
also have "... = (trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>) \<cdot>\<^sub>B (B.inv (B.\<ll> f) \<cdot>\<^sub>B B.\<ll> f)"
proof -
have "B.inv (B.\<ll> f) \<cdot>\<^sub>B B.\<ll> f = trg\<^sub>B f \<star>\<^sub>B f"
- using f ide_char B.comp_inv_arr B.inv_is_inverse by auto
+ using f ide_char\<^sub>S\<^sub>b\<^sub>C B.comp_inv_arr B.inv_is_inverse by auto
moreover have "B.dom (trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>) = trg\<^sub>B f \<star>\<^sub>B f"
proof -
have "B.dom (trg\<^sub>B ?\<mu>) = trg\<^sub>B f"
using f \<mu> B.vconn_implies_hpar(2) by force
moreover have "B.dom ?\<mu> = f"
using \<mu> by blast
ultimately show ?thesis
using B.hcomp_simps [of "trg\<^sub>B ?\<mu>" ?\<mu>]
by (metis (no_types, lifting) B.hseqI' B.ideD(1) B.src_trg
- B.trg.preserves_reflects_arr B.trg_dom f ide_char)
+ B.trg.preserves_reflects_arr B.trg_dom f ide_char\<^sub>S\<^sub>b\<^sub>C)
qed
ultimately show ?thesis
- using \<mu> \<mu>_eq B.comp_arr_dom in_hom_char by auto
+ using \<mu> \<mu>_eq B.comp_arr_dom in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
also have "... = ((trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.inv (B.\<ll> f)) \<cdot>\<^sub>B B.\<ll> f"
using B.comp_assoc by simp
also have "... = (B.inv (B.\<ll> f') \<cdot>\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.\<ll> f"
using \<mu> \<mu>_eq B.\<ll>'.naturality [of ?\<mu>] by auto
also have "... = (B.inv (B.\<ll> f') \<cdot>\<^sub>B B.\<ll> f') \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B (B.inv (B.\<ll> f) \<cdot>\<^sub>B B.\<ll> f)"
- using \<mu> \<mu>_eq arr_char arrI comp_simp B.comp_assoc by metis
+ using \<mu> \<mu>_eq arr_char\<^sub>S\<^sub>b\<^sub>C arrI comp_simp B.comp_assoc by metis
also have "... = \<nu>"
- using f f' \<nu> 2 B.comp_arr_dom B.comp_cod_arr ide_char
+ using f f' \<nu> 2 B.comp_arr_dom B.comp_cod_arr ide_char\<^sub>S\<^sub>b\<^sub>C
B.\<ll>.components_are_iso B.\<ll>_ide_simp B.comp_inv_arr'
by auto
finally show ?thesis by blast
qed
thus ?thesis
using \<mu> by auto
qed
qed
interpretation R: faithful_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> 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
+ using par inclusion arr_char\<^sub>S\<^sub>b\<^sub>C dom_simp cod_simp by fastforce
moreover have "B.R f = B.R f'"
proof -
have "\<forall>a. Arr a \<longrightarrow> B.arr a"
by (simp add: inclusion)
moreover have 1: "\<forall>a. arr a \<longrightarrow> (if arr a then hseq a (src a) else arr null)"
using R.preserves_arr by presburger
moreover have "arr f \<and> arr (src f) \<and> trg\<^sub>B (src f) = src\<^sub>B f"
by (meson 1 hcomp_def hseq_char' par)
ultimately show ?thesis
by (metis \<open>B.par f f'\<close> eq hcomp_def hseq_char' src_def)
qed
ultimately show "f = f'"
using B.R.is_faithful by blast
qed
interpretation R: full_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> R
proof
fix f f' \<nu>
assume f: "ide f" and f': "ide f'" and \<nu>: "\<guillemotleft>\<nu> : R f \<Rightarrow> R f'\<guillemotright>"
have 1: "R f = f \<star>\<^sub>B src\<^sub>B f \<and> R f' = f' \<star>\<^sub>B src\<^sub>B f'"
- using f f' hcomp_def src_def arr_char ide_char src_closed by simp
+ using f f' hcomp_def src_def arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C src_closed by simp
have 2: "\<guillemotleft>\<nu> : f \<star>\<^sub>B src\<^sub>B f \<Rightarrow>\<^sub>B f' \<star>\<^sub>B src\<^sub>B f'\<guillemotright>"
using 1 f f' \<nu> 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)
+ dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C hseq_char' arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C trg_closed null_char
+ by (simp add: arr_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C)
show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> R \<mu> = \<nu>"
proof -
let ?\<mu> = "B.\<rr> f' \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B B.inv (B.\<rr> f)"
have \<mu>: "\<guillemotleft>?\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
proof -
have "\<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
- using f f' \<nu> 2 B.\<rr>_ide_simp runit'_closed runit_closed' ide_char by auto
+ using f f' \<nu> 2 B.\<rr>_ide_simp runit'_closed runit_closed' ide_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis
- by (metis (no_types, lifting) B.arrI B.seqE \<nu> arrE arrI comp_closed f f'
- ide_char in_hom_char runit'_closed runit_closed')
+ by (metis (no_types, lifting) B.arrI B.seqE \<nu> arrE arrI\<^sub>S\<^sub>b\<^sub>C comp_closed f f'
+ ide_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C runit'_closed runit_closed')
qed
have \<mu>_eq: "?\<mu> = B.\<rr> f' \<cdot> \<nu> \<cdot> B.inv (B.\<rr> f)"
proof -
have "?\<mu> = B.\<rr> f' \<cdot> \<nu> \<cdot>\<^sub>B B.inv (B.\<rr> f)"
- using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
+ using f f' \<nu> \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C inclusion comp_char comp_closed ide_char\<^sub>S\<^sub>b\<^sub>C
runit'_closed runit_closed
by (metis (no_types, lifting) B.seqE in_homE)
also have "... = B.\<rr> f' \<cdot> \<nu> \<cdot> B.inv (B.\<rr> f)"
- using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
+ using f f' \<nu> \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C inclusion comp_char comp_closed ide_char\<^sub>S\<^sub>b\<^sub>C
runit'_closed runit_closed
- by (metis (no_types, lifting) B.arrI B.comp_in_homE in_hom_char)
+ by (metis (no_types, lifting) B.arrI B.comp_in_homE in_hom_char\<^sub>S\<^sub>b\<^sub>C)
finally show ?thesis by simp
qed
have "R ?\<mu> = \<nu>"
proof -
have "R ?\<mu> = ?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>"
- using \<mu> \<mu>_eq hcomp_def src_def inclusion arr_char src_closed by auto
+ using \<mu> \<mu>_eq hcomp_def src_def inclusion arr_char\<^sub>S\<^sub>b\<^sub>C src_closed by auto
also have "... = (?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>) \<cdot>\<^sub>B (B.inv (B.\<rr> f) \<cdot>\<^sub>B B.\<rr> f)"
proof -
have "B.inv (B.\<rr> f) \<cdot>\<^sub>B B.\<rr> f = f \<star>\<^sub>B src\<^sub>B f"
- using f ide_char B.comp_inv_arr B.inv_is_inverse by auto
+ using f ide_char\<^sub>S\<^sub>b\<^sub>C B.comp_inv_arr B.inv_is_inverse by auto
moreover have "B.dom (?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>) = f \<star>\<^sub>B src\<^sub>B f"
- using f \<mu> \<mu>_eq ide_char arr_char B.src_dom [of ?\<mu>]
+ using f \<mu> \<mu>_eq ide_char arr_char\<^sub>S\<^sub>b\<^sub>C B.src_dom [of ?\<mu>]
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 \<mu> \<mu>_eq B.comp_arr_dom in_hom_char by auto
+ using \<mu> \<mu>_eq B.comp_arr_dom in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
also have "... = ((?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.inv (B.\<rr> f)) \<cdot>\<^sub>B B.\<rr> f"
using B.comp_assoc by simp
also have "... = (B.inv (B.\<rr> f') \<cdot>\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.\<rr> f"
using \<mu> \<mu>_eq B.\<rr>'.naturality [of ?\<mu>] by auto
also have "... = (B.inv (B.\<rr> f') \<cdot>\<^sub>B B.\<rr> f') \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B (B.inv (B.\<rr> f) \<cdot>\<^sub>B B.\<rr> f)"
- using \<mu> \<mu>_eq arr_char arrI comp_simp B.comp_assoc by metis
+ using \<mu> \<mu>_eq arr_char\<^sub>S\<^sub>b\<^sub>C arrI comp_simp B.comp_assoc by metis
also have "... = \<nu>"
- using f f' \<nu> 2 B.comp_arr_dom B.comp_cod_arr ide_char
+ using f f' \<nu> 2 B.comp_arr_dom B.comp_cod_arr ide_char\<^sub>S\<^sub>b\<^sub>C
B.\<ll>.components_are_iso B.\<ll>_ide_simp B.comp_inv_arr'
by auto
finally show ?thesis by blast
qed
thus ?thesis
using \<mu> by blast
qed
qed
interpretation bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
proof
show "\<And>a. obj a \<Longrightarrow> \<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
proof (intro in_homI)
fix a
assume a: "obj a"
have 1: "Arr (\<i> 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
+ using a obj_def src_def trg_def in_hom_char\<^sub>S\<^sub>b\<^sub>C B.unit_in_hom
+ arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def B.obj_def ide_char\<^sub>S\<^sub>b\<^sub>C objE hcomp_closed
by (metis (no_types, lifting) B.\<ll>_ide_simp B.unitor_coincidence(1) inclusion lunit_closed)
show 2: "arr \<i>[a]"
- using 1 arr_char by simp
+ using 1 arr_char\<^sub>S\<^sub>b\<^sub>C by simp
show "dom \<i>[a] = a \<star> a"
- using a 2 dom_char
+ using a 2 dom_char\<^sub>S\<^sub>b\<^sub>C
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 \<i>[a] = a"
- using a 2 cod_char
+ using a 2 cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis B.obj_def' B.unit_simps(5) inclusion objE obj_simps(1)
subcategory.arrE subcategory_axioms trg_def)
qed
show "\<And>a. obj a \<Longrightarrow> iso (\<i> 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
+ using a obj_def src_def trg_def B.obj_def arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis horizontal_homs.objE horizontal_homs_axioms)
have 2: "Arr (\<i> 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
+ using a 1 obj_def src_def trg_def in_hom_char\<^sub>S\<^sub>b\<^sub>C B.unit_in_hom
+ arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def B.obj_def ide_char\<^sub>S\<^sub>b\<^sub>C objE hcomp_closed
by (metis (no_types, lifting) B.\<ll>_ide_simp B.unitor_coincidence(1) inclusion lunit_closed)
have "iso (B.\<ll> a)"
- using a 2 obj_def B.iso_unit iso_char arr_char lunit_closed lunit'_closed B.iso_lunit
+ using a 2 obj_def B.iso_unit iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C lunit_closed lunit'_closed B.iso_lunit
apply simp
by (metis (no_types, lifting) B.\<ll>.components_are_iso B.ide_src inclusion src_def)
thus "iso (\<i> a)"
- using a 2 obj_def B.iso_unit iso_char arr_char B.unitor_coincidence
+ using a 2 obj_def B.iso_unit iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C B.unitor_coincidence
apply simp
by (metis (no_types, lifting) B.\<ll>_ide_simp B.ide_src B.obj_src inclusion src_def)
qed
show "\<And>f g h k. \<lbrakk> ide f; ide g; ide h; ide k;
src f = trg g; src g = trg h; src h = trg k \<rbrakk> \<Longrightarrow>
(f \<star> \<a> g h k) \<cdot> \<a> f (g \<star> h) k \<cdot> (\<a> f g h \<star> k) =
\<a> f g (h \<star> k) \<cdot> \<a> (f \<star> 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
+ using B.pentagon VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def assoc_closed arr_char\<^sub>S\<^sub>b\<^sub>C comp_char
+ hcomp_closed comp_closed ide_char\<^sub>S\<^sub>b\<^sub>C inclusion src_def trg_def
by simp
qed
proposition is_bicategory:
shows "bicategory (\<cdot>) (\<star>) \<a> \<i> src trg"
..
lemma obj_char:
shows "obj a \<longleftrightarrow> arr a \<and> B.obj a"
proof
assume a: "obj a"
show "arr a \<and> B.obj a"
- using a obj_def B.obj_def src_def arr_char inclusion by metis
+ using a obj_def B.obj_def src_def arr_char\<^sub>S\<^sub>b\<^sub>C inclusion by metis
next
assume a: "arr a \<and> 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 = (\<lambda>f g. if arr f \<and> arr g \<and> src f = trg g then f \<star>\<^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 "\<a> f g h = \<a>\<^sub>B f g h"
- using assms VVV.arr_char VV.arr_char by auto
+ using assms VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C 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 "\<a>' f g h = B.\<a>' f g h"
proof -
have "\<a>' f g h = B.inv (\<a>\<^sub>B f g h)"
- using assms inv_char by fastforce
+ using assms inv_char\<^sub>S\<^sub>b\<^sub>C by fastforce
also have "... = B.\<a>' f g h"
- using assms ide_char src_def trg_def
- B.VVV.ide_char B.VVV.arr_char B.VV.arr_char
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
+ B.VVV.ide_char\<^sub>S\<^sub>b\<^sub>C B.VVV.arr_char\<^sub>S\<^sub>b\<^sub>C B.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
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: "\<guillemotleft>B.lunit f : trg f \<star> f \<Rightarrow> f\<guillemotright>"
proof
show 2: "arr (B.lunit f)"
- using assms arr_char lunit_closed
- by (simp add: arr_char B.\<ll>_ide_simp ide_char)
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C lunit_closed
+ by (simp add: arr_char\<^sub>S\<^sub>b\<^sub>C B.\<ll>_ide_simp ide_char\<^sub>S\<^sub>b\<^sub>C)
show "dom (B.lunit f) = trg f \<star> f"
- using assms 2 dom_char hcomp_char ide_char src_trg trg.preserves_arr trg_def
+ using assms 2 dom_char\<^sub>S\<^sub>b\<^sub>C hcomp_char ide_char\<^sub>S\<^sub>b\<^sub>C 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)
+ using assms 2 in_hom_char\<^sub>S\<^sub>b\<^sub>C
+ by (simp add: cod_simp ide_char\<^sub>S\<^sub>b\<^sub>C)
qed
show "trg f \<star> B.lunit f = (\<i>[trg f] \<star> f) \<cdot> \<a>' (trg f) (trg f) f"
proof -
have "trg f \<star> B.lunit f = trg\<^sub>B f \<star>\<^sub>B B.lunit f"
- using assms 1 arr_char hcomp_char
+ using assms 1 arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_char
by (metis (no_types, lifting) ideD(1) src_trg trg.preserves_reflects_arr
trg_def vconn_implies_hpar(2,4))
also have "... = (\<i>[trg f] \<star>\<^sub>B f) \<cdot>\<^sub>B B.\<a>' (trg f) (trg f) f"
- using assms ide_char B.lunit_char(2) trg_def by simp
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C B.lunit_char(2) trg_def by simp
also have "... = (\<i>[trg f] \<star>\<^sub>B f) \<cdot>\<^sub>B \<a>' (trg f) (trg f) f"
using assms assoc'_simp [of "trg f" "trg f" f] by simp
also have "... = (\<i>[trg f] \<star> f) \<cdot>\<^sub>B \<a>' (trg f) (trg f) f"
using assms hcomp_char by simp
also have "... = (\<i>[trg f] \<star> f) \<cdot> \<a>' (trg f) (trg f) f"
- using assms seq_char [of "\<i>[trg f] \<star> f" "\<a>' (trg f) (trg f) f"]
+ using assms seq_char\<^sub>S\<^sub>b\<^sub>C [of "\<i>[trg f] \<star> f" "\<a>' (trg f) (trg f) f"]
comp_char [of "\<i>[trg f] \<star> f" "\<a>' (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
+ using assms inv_char\<^sub>S\<^sub>b\<^sub>C [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: "\<guillemotleft>B.runit f : f \<star> src f \<Rightarrow> f\<guillemotright>"
proof
show 2: "arr (B.runit f)"
- using assms arr_char runit_closed
- by (simp add: arr_char B.\<rr>_ide_simp ide_char)
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C runit_closed
+ by (simp add: arr_char\<^sub>S\<^sub>b\<^sub>C B.\<rr>_ide_simp ide_char\<^sub>S\<^sub>b\<^sub>C)
show "dom (B.runit f) = f \<star> src f"
- using assms 2 dom_char hcomp_char
- by (metis (no_types, lifting) B.runit_simps(4) ide_char src.preserves_reflects_arr
+ using assms 2 dom_char\<^sub>S\<^sub>b\<^sub>C hcomp_char
+ by (metis (no_types, lifting) B.runit_simps(4) ide_char\<^sub>S\<^sub>b\<^sub>C 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)
+ using assms 2 in_hom_char\<^sub>S\<^sub>b\<^sub>C
+ by (simp add: cod_simp ide_char\<^sub>S\<^sub>b\<^sub>C)
qed
show "B.runit f \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a> f (src f) (src f)"
proof -
have "B.runit f \<star> src f = B.runit f \<star>\<^sub>B src\<^sub>B f"
- using assms 1 arr_char hcomp_char
+ using assms 1 arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<star>\<^sub>B \<i>[src f]) \<cdot>\<^sub>B \<a>\<^sub>B f (src f) (src f)"
- using assms ide_char B.runit_char(2) src_def by simp
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C B.runit_char(2) src_def by simp
also have "... = (f \<star>\<^sub>B \<i>[src f]) \<cdot>\<^sub>B \<a> f (src f) (src f)"
using assms assoc_simp by simp
also have "... = (f \<star> \<i>[src f]) \<cdot>\<^sub>B \<a> f (src f) (src f)"
using assms 1 hcomp_char by simp
also have "... = (f \<star> \<i>[src f]) \<cdot> \<a> f (src f) (src f)"
proof -
have "B.seq (f \<star> \<i>[src f]) (\<a> f (src f) (src f))"
- using assms seq_char [of "f \<star> \<i>[src f]" "\<a> f (src f) (src f)"] by simp
+ using assms seq_char\<^sub>S\<^sub>b\<^sub>C [of "f \<star> \<i>[src f]" "\<a> f (src f) (src f)"] by simp
thus ?thesis
using assms comp_char [of "f \<star> \<i>[src f]" "\<a> 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
+ using assms inv_char\<^sub>S\<^sub>b\<^sub>C [of "runit f"] runit_simp by fastforce
lemma comp_eqI [intro]:
assumes "seq f g" and "f = f'" and "g = g'"
shows "f \<cdot> g = f' \<cdot>\<^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 \<cdot>\<^sub>B g = f' \<cdot> 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 \<star> g = f' \<star>\<^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 \<star>\<^sub>B g = f' \<star> g'"
using assms hcomp_char not_arr_null by auto
lemma arr_compI:
assumes "seq f g"
shows "arr (f \<cdot>\<^sub>B g)"
- using assms seq_char dom_char cod_char
+ using assms seq_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) comp_simp)
lemma arr_hcompI:
assumes "hseq f g"
shows "arr (f \<star>\<^sub>B g)"
using assms hseq_char src_def trg_def hcomp_eqI' by auto
end
sublocale subbicategory \<subseteq> bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
using is_bicategory by auto
subsection "The Sub-bicategory of Endo-arrows of an Object"
text \<open>
We now consider the sub-bicategory consisting of all arrows having the same
object \<open>a\<close> 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
\<open>\<guillemotleft>\<omega> : w \<star>\<^sub>B w \<Rightarrow> w\<guillemotright>\<close> with \<open>w\<close> isomorphic to \<open>a\<close>, rather than the particular choice
\<open>\<guillemotleft>\<i>[a] : a \<star>\<^sub>B a \<Rightarrow> a\<guillemotright>\<close> made by the ambient bicategory.
\<close>
locale subbicategory_at_object =
B: bicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B +
subbicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B \<open>\<lambda>\<mu>. B.arr \<mu> \<and> src\<^sub>B \<mu> = a \<and> trg\<^sub>B \<mu> = a\<close>
for V :: "'a comp" (infixr "\<cdot>\<^sub>B" 55)
and H :: "'a comp" (infixr "\<star>\<^sub>B" 55)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
and a :: "'a"
and w :: "'a"
and \<omega> :: "'a" +
assumes obj_a: "B.obj a"
and isomorphic_a_w: "B.isomorphic a w"
and \<omega>_in_vhom: "\<guillemotleft>\<omega> : w \<star>\<^sub>B w \<Rightarrow> w\<guillemotright>"
and \<omega>_is_iso: "B.iso \<omega>"
begin
notation hcomp (infixr "\<star>" 53)
lemma arr_simps:
assumes "arr \<mu>"
shows "src \<mu> = a" and "trg \<mu> = a"
apply (metis (no_types, lifting) arrE assms src_def)
by (metis (no_types, lifting) arrE assms trg_def)
lemma \<omega>_simps [simp]:
shows "arr \<omega>"
and "src \<omega> = a" and "trg \<omega> = a"
and "dom \<omega> = w \<star>\<^sub>B w" and "cod \<omega> = w"
- using isomorphic_a_w \<omega>_in_vhom in_hom_char arr_simps by auto
+ using isomorphic_a_w \<omega>_in_vhom in_hom_char\<^sub>S\<^sub>b\<^sub>C 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 \<omega>_in_vhom ide_cod by blast
show "B.ide w"
- using w ide_char by simp
+ using w ide_char\<^sub>S\<^sub>b\<^sub>C by simp
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : a \<Rightarrow>\<^sub>B w\<guillemotright> \<and> B.iso \<phi>"
using isomorphic_a_w B.isomorphic_def by auto
show "src\<^sub>B w = a"
using obj_a w \<phi> B.src_cod by force
show "trg\<^sub>B w = a"
using obj_a w \<phi> B.src_cod by force
show "src w = a"
using \<open>src\<^sub>B w = a\<close> w ide_w src_def by simp
show "trg w = a"
using \<open>src\<^sub>B w = a\<close> w ide_w trg_def
by (simp add: \<open>trg\<^sub>B w = a\<close>)
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 \<longleftrightarrow> VV.arr f"
- using inclusion VxV.arr_char VV.arr_char arr_char src_def trg_def
+ using inclusion VxV.arr_char VV.arr_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def
by auto
lemma VxV_comp_eq_VV_comp:
shows "VxV.comp = VV.comp"
proof -
have "\<And>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 \<longleftrightarrow> VVV.arr f"
- using VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char
+ using VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def inclusion arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
lemma VxVxV_comp_eq_VVV_comp:
shows "VxVxV.comp = VVV.comp"
proof -
have "\<And>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: "\<not> 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 "\<not> VVV.seq f g"
- using 1 VVV.seq_char by blast
+ using 1 VVV.seq_char\<^sub>S\<^sub>b\<^sub>C 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) \<and> B.arr (fst (snd f)) \<and> B.arr (snd (snd f)) \<and>
src\<^sub>B (fst f) = a \<and> src\<^sub>B (fst (snd f)) = a \<and> src\<^sub>B (snd (snd f)) = a \<and>
trg\<^sub>B (fst f) = a \<and> trg\<^sub>B (fst (snd f)) = a \<and> trg\<^sub>B (snd (snd f)) = a"
- using 1 VxVxV.seq_char VxV.seq_char arr_char by blast
+ using 1 VxVxV.seq_char VxV.seq_char arr_char\<^sub>S\<^sub>b\<^sub>C by blast
have 3: "B.arr (fst g) \<and> B.arr (fst (snd g)) \<and> B.arr (snd (snd g)) \<and>
src\<^sub>B (fst g) = a \<and> src\<^sub>B (fst (snd g)) = a \<and> src\<^sub>B (snd (snd g)) = a \<and>
trg\<^sub>B (fst g) = a \<and> trg\<^sub>B (fst (snd g)) = a \<and> trg\<^sub>B (snd (snd g)) = a"
- using 1 VxVxV.seq_char VxV.seq_char arr_char by blast
+ using 1 VxVxV.seq_char VxV.seq_char arr_char\<^sub>S\<^sub>b\<^sub>C by blast
have 4: "B.seq (fst f) (fst g) \<and> B.seq (fst (snd f)) (fst (snd g)) \<and>
B.seq (snd (snd f)) (snd (snd g))"
- using 1 VxVxV.seq_char VxV.seq_char seq_char by blast
+ using 1 VxVxV.seq_char VxV.seq_char seq_char\<^sub>S\<^sub>b\<^sub>C by blast
have 5: "VxVxV.comp f g =
(fst f \<cdot> fst g, fst (snd f) \<cdot> fst (snd g), snd (snd f) \<cdot> snd (snd g))"
- using 1 2 3 4 VxVxV.seqE VxVxV.comp_char VxV.comp_char seq_char arr_char
+ using 1 2 3 4 VxVxV.seqE VxVxV.comp_char VxV.comp_char seq_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting))
also have "... = VVV.comp f g"
- using 1 VVV.comp_char VVV.arr_char VV.arr_char
+ using 1 VVV.comp_char VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply simp
- using 2 3 5 arrI arr_simps(1) arr_simps(2) by presburger
+ using 2 3 5 arrI\<^sub>S\<^sub>b\<^sub>C 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 \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
using H.functor_axioms hcomp_def VxV_comp_eq_VV_comp by simp
interpretation H: binary_endofunctor \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> ..
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
+ using HoHV_def H.ToTC_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def inclusion arr_char\<^sub>S\<^sub>b\<^sub>C
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
+ using HoVH_def H.ToCT_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def inclusion arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
interpretation ToTC: "functor" VxVxV.comp \<open>(\<cdot>)\<close> H.ToTC
using HoHV_eq_ToTC VxVxV_comp_eq_VVV_comp HoHV.functor_axioms by simp
interpretation ToCT: "functor" VxVxV.comp \<open>(\<cdot>)\<close> H.ToCT
using HoVH_eq_ToCT VxVxV_comp_eq_VVV_comp HoVH.functor_axioms by simp
interpretation \<alpha>: natural_isomorphism VxVxV.comp \<open>(\<cdot>)\<close> H.ToTC H.ToCT \<alpha>
unfolding \<alpha>_def
using \<alpha>.natural_isomorphism_axioms HoHV_eq_ToTC HoVH_eq_ToCT \<alpha>_def
VxVxV_comp_eq_VVV_comp
by simp
interpretation L: endofunctor \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (w, f) \<star> snd (w, f)\<close>
proof
fix f
show "\<not> arr f \<Longrightarrow> fst (w, f) \<star> snd (w, f) = null"
- using arr_char hcomp_def by auto
+ using arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<omega>_in_vhom cod_char by simp
+ using f hseq_char arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def \<omega>_in_vhom cod_char\<^sub>S\<^sub>b\<^sub>C by simp
show "dom (fst (w, f) \<star> snd (w, f)) = fst (w, dom f) \<star> snd (w, dom f)"
- using f arr_char hcomp_def dom_simp by simp
+ using f arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def dom_simp by simp
show "cod (fst (w, f) \<star> snd (w, f)) = fst (w, cod f) \<star> snd (w, cod f)"
- using f arr_char hcomp_def cod_simp by simp
+ using f arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def cod_simp by simp
next
fix f g
assume fg: "seq g f"
show "fst (w, g \<cdot> f) \<star> snd (w, g \<cdot> f) = (fst (w, g) \<star> snd (w, g)) \<cdot> (fst (w, f) \<star> snd (w, f))"
by (simp add: fg whisker_left)
qed
interpretation L': equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (w, f) \<star> snd (w, f)\<close>
proof -
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi>"
using isomorphic_a_w B.isomorphic_symmetric by force
have "\<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright>"
- using \<phi> in_hom_char
+ using \<phi> in_hom_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) B.in_homE B.src_cod B.src_src B.trg_cod B.trg_trg
- \<omega>_in_vhom arr_char arr_cod cod_simp)
+ \<omega>_in_vhom arr_char\<^sub>S\<^sub>b\<^sub>C arr_cod cod_simp)
hence \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi> \<and> \<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright> \<and> iso \<phi>"
- using \<phi> iso_char arr_char by auto
+ using \<phi> iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by auto
interpret \<l>: natural_isomorphism \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close>
\<open>\<lambda>f. fst (w, f) \<star> snd (w, f)\<close> map \<open>\<lambda>f. \<ll> f \<cdot> (\<phi> \<star> dom f)\<close>
proof
fix \<mu>
show "\<not> arr \<mu> \<Longrightarrow> \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>) = null"
- using \<phi> arr_char dom_char ext
+ using \<phi> arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C ext
apply simp
using comp_null(2) hcomp_def by fastforce
assume \<mu>: "arr \<mu>"
have 0: "in_hhom (dom \<mu>) a a"
- using \<mu> arr_char src_dom trg_dom src_def trg_def dom_simp by simp
+ using \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C src_dom trg_dom src_def trg_def dom_simp by simp
have 1: "in_hhom \<phi> a a"
- using \<phi> arr_char src_dom trg_dom src_def trg_def by auto
+ using \<phi> arr_char\<^sub>S\<^sub>b\<^sub>C src_dom trg_dom src_def trg_def by auto
have 2: "hseq \<phi> (B.dom \<mu>)"
using \<mu> 0 1 dom_simp by (intro hseqI) auto
have 3: "seq (\<ll> \<mu>) (\<phi> \<star> dom \<mu>)"
proof (intro seqI')
show "\<guillemotleft>\<phi> \<star> dom \<mu> : w \<star> dom \<mu> \<Rightarrow> a \<star> dom \<mu>\<guillemotright>"
by (metis (no_types, lifting) 0 \<mu> \<phi> hcomp_in_vhom ide_dom ide_in_hom(2)
in_hhom_def w_simps(3))
show "\<guillemotleft>\<ll> \<mu> : a \<star> dom \<mu> \<Rightarrow> cod \<mu>\<guillemotright>"
using \<mu> 2 \<ll>.preserves_hom [of \<mu> "dom \<mu>" "cod \<mu>"] arr_simps(2) arr_cod by fastforce
qed
show "dom (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = fst (w, dom \<mu>) \<star> snd (w, dom \<mu>)"
proof -
have "dom (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = dom \<phi> \<star> dom \<mu>"
using \<mu> 3 hcomp_simps(3) dom_comp dom_dom
apply (elim seqE) by auto
also have "... = fst (w, dom \<mu>) \<star> snd (w, dom \<mu>)"
using \<omega>_in_vhom \<phi>
by (metis (no_types, lifting) in_homE prod.sel(1) prod.sel(2))
finally show ?thesis by simp
qed
show "cod (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = map (cod \<mu>)"
proof -
have "seq (\<ll> \<mu>) (\<phi> \<star> dom \<mu>)"
using 3 by simp
hence "cod (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = cod (\<ll> \<mu>)"
using cod_comp by blast
also have "... = map (cod \<mu>)"
using \<mu> by blast
finally show ?thesis by blast
qed
show "map \<mu> \<cdot> \<ll> (dom \<mu>) \<cdot> (\<phi> \<star> dom (dom \<mu>)) = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "map \<mu> \<cdot> \<ll> (dom \<mu>) \<cdot> (\<phi> \<star> dom (dom \<mu>)) = (map \<mu> \<cdot> \<ll> (dom \<mu>)) \<cdot> (\<phi> \<star> dom \<mu>)"
using \<mu> comp_assoc by simp
also have "... = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
using \<mu> \<phi> \<ll>.is_natural_1 by auto
finally show ?thesis by blast
qed
show "(\<ll> (cod \<mu>) \<cdot> (\<phi> \<star> dom (cod \<mu>))) \<cdot> (fst (w, \<mu>) \<star> snd (w, \<mu>)) = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "(\<ll> (cod \<mu>) \<cdot> (\<phi> \<star> dom (cod \<mu>))) \<cdot> (fst (w, \<mu>) \<star> snd (w, \<mu>)) =
(\<ll> (cod \<mu>) \<cdot> (\<phi> \<star> B.cod \<mu>)) \<cdot> (w \<star> \<mu>)"
- using \<mu> \<phi> dom_char arr_char \<omega>_in_vhom cod_simp by simp
+ using \<mu> \<phi> dom_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C \<omega>_in_vhom cod_simp by simp
also have "... = \<ll> (cod \<mu>) \<cdot> (\<phi> \<cdot> w \<star> B.cod \<mu> \<cdot> \<mu>)"
proof -
have "seq \<phi> w"
using \<phi> \<omega>_in_vhom w_simps(1) by blast
moreover have 2: "seq (B.cod \<mu>) \<mu>"
- using \<mu> seq_char cod_simp by (simp add: comp_cod_arr)
+ using \<mu> seq_char\<^sub>S\<^sub>b\<^sub>C cod_simp by (simp add: comp_cod_arr)
moreover have "src \<phi> = trg (B.cod \<mu>)"
using \<mu> \<phi> 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 "... = \<ll> (cod \<mu>) \<cdot> (\<phi> \<star> \<mu>)"
using \<mu> \<phi> \<omega>_in_vhom comp_arr_dom comp_cod_arr cod_simp
apply (elim conjE in_homE) by auto
also have "... = (\<ll> (cod \<mu>) \<cdot> (cod \<phi> \<star> \<mu>)) \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have 1: "seq (cod \<phi>) \<phi>"
using \<phi> arr_cod_iff_arr dom_cod iso_is_arr seqI by presburger
moreover have 2: "seq \<mu> (dom \<mu>)"
using \<mu> by (simp add: comp_arr_dom)
moreover have "src (cod \<phi>) = trg \<mu>"
using \<mu> \<phi> arr_cod arr_simps(1-2) iso_is_arr by auto
ultimately show ?thesis
using 1 2 interchange [of "cod \<phi>" \<phi> \<mu> "dom \<mu>"] comp_arr_dom comp_cod_arr
comp_assoc by fastforce
qed
also have "... = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "L \<mu> = cod \<phi> \<star> \<mu>"
using \<mu> \<phi> arr_simps(2) in_homE by auto
hence "\<ll> (cod \<mu>) \<cdot> (cod \<phi> \<star> \<mu>) = \<ll> \<mu>"
using \<mu> \<ll>.is_natural_2 [of \<mu>] by simp
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
next
show "\<And>f. ide f \<Longrightarrow> iso (\<ll> f \<cdot> (\<phi> \<star> dom f))"
proof -
fix f
assume f: "ide f"
have "iso (\<ll> f)"
using f iso_lunit by simp
moreover have "iso (\<phi> \<star> dom f)"
- using \<phi> f src_def trg_def ide_char arr_char
+ using \<phi> f src_def trg_def ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
apply (intro iso_hcomp, simp_all)
by (metis (no_types, lifting) in_homE)
moreover have "seq (\<ll> f) (\<phi> \<star> dom f)"
proof (intro seqI')
show " \<guillemotleft>\<ll> f : a \<star> f \<Rightarrow> f\<guillemotright>"
- using f lunit_in_hom(2) \<ll>_ide_simp ide_char arr_char trg_def by simp
+ using f lunit_in_hom(2) \<ll>_ide_simp ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C trg_def by simp
show "\<guillemotleft>\<phi> \<star> dom f : w \<star> f \<Rightarrow> a \<star> f\<guillemotright>"
- using \<phi> f ide_char arr_char hcomp_def src_def trg_def obj_a ide_in_hom
- in_hom_char
+ using \<phi> f ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def src_def trg_def obj_a ide_in_hom
+ in_hom_char\<^sub>S\<^sub>b\<^sub>C
by (intro hcomp_in_vhom, auto)
qed
ultimately show "iso (\<ll> f \<cdot> (\<phi> \<star> dom f))"
using isos_compose by simp
qed
qed
show "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (w, f) \<star> snd (w, f))"
using \<l>.natural_isomorphism_axioms L.isomorphic_to_identity_is_equivalence by simp
qed
interpretation L: equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (cod \<omega>, f) \<star> snd (cod \<omega>, f)\<close>
proof -
have "(\<lambda>f. fst (cod \<omega>, f) \<star> snd (cod \<omega>, f)) = (\<lambda>f. fst (w, f) \<star> snd (w, f))"
using \<omega>_in_vhom by simp
thus "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (cod \<omega>, f) \<star> snd (cod \<omega>, f))"
using L'.equivalence_functor_axioms by simp
qed
interpretation R: endofunctor \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (f, w) \<star> snd (f, w)\<close>
proof
fix f
show "\<not> arr f \<Longrightarrow> fst (f, w) \<star> snd (f, w) = null"
- using arr_char hcomp_def by auto
+ using arr_char\<^sub>S\<^sub>b\<^sub>C 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 \<omega>_in_vhom cod_char isomorphic_a_w
- B.isomorphic_def in_hom_char
+ using f hseq_char arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def \<omega>_in_vhom cod_char\<^sub>S\<^sub>b\<^sub>C isomorphic_a_w
+ B.isomorphic_def in_hom_char\<^sub>S\<^sub>b\<^sub>C
by simp
show "dom (fst (f, w) \<star> snd (f, w)) = fst (dom f, w) \<star> snd (dom f, w)"
- using f arr_char dom_char cod_char hcomp_def \<omega>_in_vhom by simp
+ using f arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C hcomp_def \<omega>_in_vhom by simp
show "cod (fst (f, w) \<star> snd (f, w)) = fst (cod f, w) \<star> snd (cod f, w)"
- using f arr_char dom_char cod_char hcomp_def \<omega>_in_vhom by simp
+ using f arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C hcomp_def \<omega>_in_vhom by simp
next
fix f g
assume fg: "seq g f"
have 1: "a \<cdot>\<^sub>B a = a"
using obj_a by auto
show "fst (g \<cdot> f, w) \<star> snd (g \<cdot> f, w) = (fst (g, w) \<star> snd (g, w)) \<cdot> (fst (f, w) \<star> snd (f, w))"
by (simp add: fg whisker_right)
qed
interpretation R': equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (f, w) \<star> snd (f, w)\<close>
proof -
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi>"
using isomorphic_a_w B.isomorphic_symmetric by force
have "\<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright>"
- using \<phi> in_hom_char
+ using \<phi> in_hom_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) B.in_homE B.src_cod B.src_src B.trg_cod B.trg_trg
- \<omega>_in_vhom arr_char arr_cod cod_simp)
+ \<omega>_in_vhom arr_char\<^sub>S\<^sub>b\<^sub>C arr_cod cod_simp)
hence \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi> \<and> \<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright> \<and> iso \<phi>"
- using \<phi> iso_char arr_char by auto
+ using \<phi> iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by auto
interpret \<r>: natural_isomorphism comp comp
\<open>\<lambda>f. fst (f, w) \<star> snd (f, w)\<close> map \<open>\<lambda>f. \<rr> f \<cdot> (dom f \<star> \<phi>)\<close>
proof
fix \<mu>
show "\<not> arr \<mu> \<Longrightarrow> \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>) = null"
- using \<phi> arr_char dom_char ext
+ using \<phi> arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C ext
apply simp
using comp_null(2) hcomp_def by fastforce
assume \<mu>: "arr \<mu>"
have 0: "in_hhom (dom \<mu>) a a"
- using \<mu> arr_char src_dom trg_dom src_def trg_def dom_simp by simp
+ using \<mu> arr_char\<^sub>S\<^sub>b\<^sub>C src_dom trg_dom src_def trg_def dom_simp by simp
have 1: "in_hhom \<phi> a a"
- using \<phi> arr_char src_dom trg_dom src_def trg_def by auto
+ using \<phi> arr_char\<^sub>S\<^sub>b\<^sub>C src_dom trg_dom src_def trg_def by auto
have 2: "hseq (B.dom \<mu>) \<phi>"
using \<mu> 0 1 dom_simp hseqI by auto
have 3: "seq (\<rr> \<mu>) (dom \<mu> \<star> \<phi>)"
proof (intro seqI')
show "\<guillemotleft>dom \<mu> \<star> \<phi> : dom \<mu> \<star> w \<Rightarrow> dom \<mu> \<star> a\<guillemotright>"
by (metis (no_types, lifting) "0" "1" \<mu> \<phi> hcomp_in_vhom hseqI hseq_char
ide_dom ide_in_hom(2) vconn_implies_hpar(2))
show "\<guillemotleft>\<rr> \<mu> : dom \<mu> \<star> a \<Rightarrow> cod \<mu>\<guillemotright>"
using \<mu> 2 \<rr>.preserves_hom [of \<mu> "dom \<mu>" "cod \<mu>"] arr_simps(2) arr_cod
dom_simp cod_simp
by fastforce
qed
show "dom (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = fst (dom \<mu>, w) \<star> snd (dom \<mu>, w)"
proof -
have "dom (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = dom \<mu> \<star> dom \<phi>"
using \<mu> 3 hcomp_simps(3) dom_comp dom_dom
apply (elim seqE) by auto
also have "... = fst (dom \<mu>, w) \<star> snd (dom \<mu>, w)"
using \<omega>_in_vhom \<phi>
by (metis (no_types, lifting) in_homE prod.sel(1) prod.sel(2))
finally show ?thesis by simp
qed
show "cod (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = map (cod \<mu>)"
proof -
have "seq (\<rr> \<mu>) (dom \<mu> \<star> \<phi>)"
using 3 by simp
hence "cod (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = cod (\<rr> \<mu>)"
using cod_comp by blast
also have "... = map (cod \<mu>)"
using \<mu> by blast
finally show ?thesis by blast
qed
show "map \<mu> \<cdot> \<rr> (dom \<mu>) \<cdot> (dom (dom \<mu>) \<star> \<phi>) = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "map \<mu> \<cdot> \<rr> (dom \<mu>) \<cdot> (dom (dom \<mu>) \<star> \<phi>) =
(map \<mu> \<cdot> \<rr> (dom \<mu>)) \<cdot> (dom (dom \<mu>) \<star> \<phi>)"
using comp_assoc by simp
also have "... = (map \<mu> \<cdot> \<rr> (dom \<mu>)) \<cdot> (dom \<mu> \<star> \<phi>)"
using \<mu> dom_dom by simp
also have "... = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
using \<mu> \<phi> \<rr>.is_natural_1 by auto
finally show ?thesis by blast
qed
show "(\<rr> (cod \<mu>) \<cdot> (dom (cod \<mu>) \<star> \<phi>)) \<cdot> (fst (\<mu>, w) \<star> snd (\<mu>, w)) = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "(\<rr> (cod \<mu>) \<cdot> (dom (cod \<mu>) \<star> \<phi>)) \<cdot> (fst (\<mu>, w) \<star> snd (\<mu>, w)) =
(\<rr> (cod \<mu>) \<cdot> (B.cod \<mu> \<star> \<phi>)) \<cdot> (\<mu> \<star> w)"
- using \<mu> \<phi> dom_char arr_char \<omega>_in_vhom cod_simp by simp
+ using \<mu> \<phi> dom_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C \<omega>_in_vhom cod_simp by simp
also have "... = \<rr> (cod \<mu>) \<cdot> (B.cod \<mu> \<cdot> \<mu> \<star> \<phi> \<cdot> w)"
proof -
have 2: "seq \<phi> w"
using \<phi> \<omega>_in_vhom w_simps(1) by blast
moreover have "seq (B.cod \<mu>) \<mu>"
- using \<mu> seq_char cod_simp by (simp add: comp_cod_arr)
+ using \<mu> seq_char\<^sub>S\<^sub>b\<^sub>C cod_simp by (simp add: comp_cod_arr)
moreover have "src (B.cod \<mu>) = trg \<phi>"
using \<mu> \<phi> 2
- using arr_simps(1) calculation(2) seq_char vconn_implies_hpar(2) by force
+ using arr_simps(1) calculation(2) seq_char\<^sub>S\<^sub>b\<^sub>C vconn_implies_hpar(2) by force
ultimately show ?thesis
using interchange comp_assoc by simp
qed
also have "... = \<rr> (cod \<mu>) \<cdot> (\<mu> \<star> \<phi>)"
using \<mu> \<phi> \<omega>_in_vhom comp_arr_dom comp_cod_arr cod_simp
apply (elim conjE in_homE) by auto
also have "... = (\<rr> (cod \<mu>) \<cdot> (\<mu> \<star> cod \<phi>)) \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "(\<mu> \<star> cod \<phi>) \<cdot> (dom \<mu> \<star> \<phi>) = \<mu> \<star> \<phi>"
proof -
have "seq \<mu> (dom \<mu>)"
using \<mu> by (simp add: comp_arr_dom)
moreover have "seq (cod \<phi>) \<phi>"
using \<phi> iso_is_arr arr_cod dom_cod by auto
moreover have "src \<mu> = trg (cod \<phi>)"
using \<mu> \<phi> 2
by (metis (no_types, lifting) arr_simps(1) arr_simps(2) calculation(2) seqE)
ultimately show ?thesis
using \<mu> \<phi> iso_is_arr comp_arr_dom comp_cod_arr
interchange [of \<mu> "dom \<mu>" "cod \<phi>" \<phi>]
by simp
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "\<mu> \<star> cod \<phi> = R \<mu>"
using \<mu> \<phi> arr_simps(1) in_homE by auto
hence "\<rr> (cod \<mu>) \<cdot> (\<mu> \<star> cod \<phi>) = \<rr> \<mu>"
using \<mu> \<phi> \<rr>.is_natural_2 by simp
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
next
show "\<And>f. ide f \<Longrightarrow> iso (\<rr> f \<cdot> (dom f \<star> \<phi>))"
proof -
fix f
assume f: "ide f"
have 1: "iso (\<rr> f)"
using f iso_lunit by simp
moreover have 2: "iso (dom f \<star> \<phi>)"
- using \<phi> f src_def trg_def ide_char arr_char
+ using \<phi> f src_def trg_def ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
apply (intro iso_hcomp, simp_all)
by (metis (no_types, lifting) in_homE)
moreover have "seq (\<rr> f) (dom f \<star> \<phi>)"
proof (intro seqI')
show "\<guillemotleft>\<rr> f : f \<star> a \<Rightarrow> f\<guillemotright>"
- using f runit_in_hom(2) \<rr>_ide_simp ide_char arr_char src_def by simp
+ using f runit_in_hom(2) \<rr>_ide_simp ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C src_def by simp
show "\<guillemotleft>dom f \<star> \<phi> : f \<star> w \<Rightarrow> f \<star> a\<guillemotright>"
- using \<phi> f ide_char arr_char hcomp_def src_def trg_def obj_a ide_in_hom
- in_hom_char
+ using \<phi> f ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C hcomp_def src_def trg_def obj_a ide_in_hom
+ in_hom_char\<^sub>S\<^sub>b\<^sub>C
by (intro hcomp_in_vhom, auto)
qed
ultimately show "iso (\<rr> f \<cdot> (dom f \<star> \<phi>))"
using isos_compose by blast
qed
qed
show "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (f, w) \<star> snd (f, w))"
using \<r>.natural_isomorphism_axioms R.isomorphic_to_identity_is_equivalence by simp
qed
interpretation R: equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (f, cod \<omega>) \<star> snd (f, cod \<omega>)\<close>
proof -
have "(\<lambda>f. fst (f, cod \<omega>) \<star> snd (f, cod \<omega>)) = (\<lambda>f. fst (f, w) \<star> snd (f, w))"
using \<omega>_in_vhom by simp
thus "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (f, cod \<omega>) \<star> snd (f, cod \<omega>))"
using R'.equivalence_functor_axioms by simp
qed
interpretation M: monoidal_category \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> \<alpha> \<omega>
proof
show "\<guillemotleft>\<omega> : fst (cod \<omega>, cod \<omega>) \<star> snd (cod \<omega>, cod \<omega>) \<Rightarrow> cod \<omega>\<guillemotright>"
- using \<omega>_in_vhom hcomp_def arr_char by auto
+ using \<omega>_in_vhom hcomp_def arr_char\<^sub>S\<^sub>b\<^sub>C by auto
show "iso \<omega>"
- using \<omega>_is_iso iso_char arr_char inv_char \<omega>_in_vhom by auto
+ using \<omega>_is_iso iso_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C inv_char\<^sub>S\<^sub>b\<^sub>C \<omega>_in_vhom by auto
show "\<And>f g h k. \<lbrakk> ide f; ide g; ide h; ide k \<rbrakk> \<Longrightarrow>
(fst (f, \<alpha> (g, h, k)) \<star> snd (f, \<alpha> (g, h, k))) \<cdot>
\<alpha> (f, hcomp (fst (g, h)) (snd (g, h)), k) \<cdot>
(fst (\<alpha> (f, g, h), k) \<star> snd (\<alpha> (f, g, h), k)) =
\<alpha> (f, g, fst (h, k) \<star> snd (h, k)) \<cdot> \<alpha> (fst (f, g) \<star> 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) \<and> VVV.arr (g, h, k)"
- using f g h k VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char
+ using f g h k VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
by simp
have 2: "VVV.arr (f, g \<star> h, k)"
- using f g h k 1 HoHV_def VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char
+ using f g h k 1 HoHV_def VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp hseqI'
by auto
have 3: "VVV.arr (f, g, h \<star> k)"
- using f g h k 1 VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char
+ using f g h k 1 VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C
VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp H.preserves_reflects_arr hseqI'
by auto
have 4: "VVV.arr (f \<star> g, h, k)"
- using f g h k VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char hseq_char
+ using f g h k VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C src_def trg_def ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C hseq_char
VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp
by force
have "(fst (f, \<alpha> (g, h, k)) \<star> snd (f, \<alpha> (g, h, k))) \<cdot>
\<alpha> (f, fst (g, h) \<star> snd (g, h), k) \<cdot>
(fst (\<alpha> (f, g, h), k) \<star> snd (\<alpha> (f, g, h), k)) =
(f \<star> \<a>\<^sub>B[g, h, k]) \<cdot> \<a>\<^sub>B[f, g \<star> h, k] \<cdot> (\<a>\<^sub>B[f, g, h] \<star> k)"
unfolding \<alpha>_def by (simp add: 1 2)
also have "... = (f \<star>\<^sub>B \<a>\<^sub>B g h k) \<cdot> \<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot> (\<a>\<^sub>B f g h \<star>\<^sub>B k)"
unfolding hcomp_def
- using f g h k src_def trg_def arr_char
- using assoc_closed ide_char by auto
+ using f g h k src_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C
+ using assoc_closed ide_char\<^sub>S\<^sub>b\<^sub>C by auto
also have "... = (f \<star>\<^sub>B \<a>\<^sub>B g h k) \<cdot>\<^sub>B \<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B (\<a>\<^sub>B f g h \<star>\<^sub>B k)"
proof -
have "arr (f \<star>\<^sub>B \<a>\<^sub>B g h k)"
- using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f (g \<star>\<^sub>B h) k)"
- using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f g h \<star>\<^sub>B k)"
- using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B (\<a>\<^sub>B f g h \<star>\<^sub>B k))"
- unfolding arr_char
+ unfolding arr_char\<^sub>S\<^sub>b\<^sub>C
apply (intro conjI)
- using ide_char arr_char assoc_closed f g h hcomp_closed k B.HoHV_def B.HoVH_def
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C 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 (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B \<a>\<^sub>B f g h \<star>\<^sub>B k)"
- using f g h k ide_char arr_char B.HoHV_def B.HoVH_def
+ using f g h k ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C B.HoHV_def B.HoVH_def
apply (intro B.seqI)
by auto
show "src\<^sub>B (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B \<a>\<^sub>B f g h \<star>\<^sub>B k) = a"
- using 1 f g h k arr_char B.src_vcomp B.vseq_implies_hpar(1) by fastforce
+ using 1 f g h k arr_char\<^sub>S\<^sub>b\<^sub>C B.src_vcomp B.vseq_implies_hpar(1) by fastforce
show "trg\<^sub>B (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B \<a>\<^sub>B f g h \<star>\<^sub>B k) = a"
- using "1" arr_char calculation(2-3) by auto
+ using "1" arr_char\<^sub>S\<^sub>b\<^sub>C calculation(2-3) by auto
qed
ultimately show ?thesis
using B.ext comp_char by (metis (no_types, lifting))
qed
also have "... = \<a>\<^sub>B f g (h \<star>\<^sub>B k) \<cdot>\<^sub>B \<a>\<^sub>B (f \<star>\<^sub>B g) h k"
- using f g h k src_def trg_def arr_char ide_char B.pentagon
+ using f g h k src_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C B.pentagon
using "4" \<alpha>_def hcomp_def by auto
also have "... = \<a>\<^sub>B f g (h \<star>\<^sub>B k) \<cdot> \<a>\<^sub>B (f \<star>\<^sub>B g) h k"
proof -
have "arr (\<a>\<^sub>B (f \<star>\<^sub>B g) h k)"
- using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f g (h \<star>\<^sub>B k))"
- using ide_char arr_char assoc_closed f g h hcomp_closed k by fastforce
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C assoc_closed f g h hcomp_closed k by fastforce
ultimately show ?thesis
using B.ext comp_char by auto
qed
also have "... = \<a>\<^sub>B[f, g, fst (h, k) \<star> snd (h, k)] \<cdot> \<a>\<^sub>B[fst (f, g) \<star> snd (f, g), h, k]"
unfolding hcomp_def
- using f g h k src_def trg_def arr_char ide_char by simp
+ using f g h k src_def trg_def arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C by simp
also have "... = \<alpha> (f, g, fst (h, k) \<star> snd (h, k)) \<cdot> \<alpha> (fst (f, g) \<star> snd (f, g), h, k)"
unfolding \<alpha>_def using 1 2 3 4 by simp
finally show "(fst (f, \<alpha> (g, h, k)) \<star> snd (f, \<alpha> (g, h, k))) \<cdot>
\<alpha> (f, fst (g, h) \<star> snd (g, h), k) \<cdot>
(fst (\<alpha> (f, g, h), k) \<star> snd (\<alpha> (f, g, h), k)) =
\<alpha> (f, g, fst (h, k) \<star> snd (h, k)) \<cdot> \<alpha> (fst (f, g) \<star> snd (f, g), h, k)"
by simp
qed
qed
proposition is_monoidal_category:
shows "monoidal_category (\<cdot>) (\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>) \<alpha> \<omega>"
..
end
text \<open>
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 \<open>a\<close> is additionally equipped with an arbitrarily chosen
unit isomorphism \<open>\<guillemotleft>\<iota> : a \<star> a \<Rightarrow> a\<guillemotright>\<close>. For any \<open>(a, \<iota>)\<close> and \<open>(a', \<iota>')\<close>,
where \<open>a\<close> and \<open>a'\<close> are isomorphic to the same object, there exists a unique isomorphism
\<open>\<guillemotleft>\<psi>: a \<Rightarrow> a'\<guillemotright>\<close> that is compatible with the chosen unit isomorphisms \<open>\<iota>\<close> and \<open>\<iota>'\<close>.
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 \<open>a\<close> is an object
in a bicategory, then the sub-bicategory consisting of all \<open>\<mu>\<close> such that
\<open>src \<mu> = a = trg \<mu>\<close>, 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.
\<close>
lemma (in bicategory) unit_unique_upto_unique_iso:
assumes "obj a"
and "isomorphic a w"
and "\<guillemotleft>\<omega> : w \<star> w \<Rightarrow> w\<guillemotright>"
and "iso \<omega>"
shows "\<exists>!\<psi>. \<guillemotleft>\<psi> : a \<Rightarrow> w\<guillemotright> \<and> iso \<psi> \<and> \<psi> \<cdot> \<i>[a] = \<omega> \<cdot> (\<psi> \<star> \<psi>)"
proof -
have \<omega>_in_hhom: "\<guillemotleft>\<omega> : a \<rightarrow> a\<guillemotright>"
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 \<a> \<i> src trg \<open>\<lambda>\<mu>. arr \<mu> \<and> src \<mu> = a \<and> trg \<mu> = a\<close>
- using assms iso_unit in_homE isoE isomorphicE VVV.arr_char VV.arr_char
+ using assms iso_unit in_homE isoE isomorphicE VVV.arr_char\<^sub>S\<^sub>b\<^sub>C VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply unfold_locales
apply auto[7]
proof
fix f g h
assume f: "(arr f \<and> src f = a \<and> trg f = a) \<and> ide f"
and g: "(arr g \<and> src g = a \<and> trg g = a) \<and> ide g"
and h: "(arr h \<and> src h = a \<and> trg h = a) \<and> ide h"
and fg: "src f = trg g" and gh: "src g = trg h"
show "arr (\<a>[f, g, h])"
using assms f g h fg gh by auto
show "src (\<a>[f, g, h]) = a \<and> trg (\<a>[f, g, h]) = a"
using assms f g h fg gh by auto
show "arr (inv (\<a>[f, g, h])) \<and> src (inv (\<a>[f, g, h])) = a \<and> trg (inv (\<a>[f, g, h])) = a"
using assms f g h fg gh \<alpha>.preserves_hom src_dom trg_dom by simp
next
fix f
assume f: "arr f \<and> src f = a \<and> trg f = a"
assume ide_left: "ide f"
show "arr (\<ll> f) \<and> src (\<ll> f) = a \<and> trg (\<ll> f) = a"
using f assms(1) \<ll>.preserves_hom src_cod [of "\<ll> f"] trg_cod [of "\<ll> f"] by simp
show "arr (inv (\<ll> f)) \<and> src (inv (\<ll> f)) = a \<and> trg (inv (\<ll> f)) = a"
using f ide_left assms(1) \<ll>'.preserves_hom src_dom [of "\<ll>'.map f"] trg_dom [of "\<ll>'.map f"]
by simp
show "arr (\<rr> f) \<and> src (\<rr> f) = a \<and> trg (\<rr> f) = a"
using f assms(1) \<rr>.preserves_hom src_cod [of "\<rr> f"] trg_cod [of "\<rr> f"] by simp
show "arr (inv (\<rr> f)) \<and> src (inv (\<rr> f)) = a \<and> trg (inv (\<rr> f)) = a"
using f ide_left assms(1) \<rr>'.preserves_hom src_dom [of "\<rr>'.map f"] trg_dom [of "\<rr>'.map f"]
by simp
qed
interpret S: subbicategory_at_object V H \<a> \<i> src trg a a \<open>\<i>[a]\<close>
proof
show "obj a" by fact
show "isomorphic a a"
using assms(1) isomorphic_reflexive by blast
show "S.in_hom \<i>[a] (a \<star> a) a"
- using S.arr_char S.in_hom_char assms(1) by fastforce
+ using S.arr_char\<^sub>S\<^sub>b\<^sub>C S.in_hom_char\<^sub>S\<^sub>b\<^sub>C assms(1) by fastforce
show "iso \<i>[a]"
using assms iso_unit by simp
qed
interpret S\<^sub>\<omega>: subbicategory_at_object V H \<a> \<i> src trg a w \<omega>
proof
show "obj a" by fact
show "iso \<omega>" by fact
show "isomorphic a w"
using assms by simp
show "S.in_hom \<omega> (w \<star> w) w"
- using assms S.arr_char S.dom_char S.cod_char \<omega>_in_hhom
+ using assms S.arr_char\<^sub>S\<^sub>b\<^sub>C S.dom_char\<^sub>S\<^sub>b\<^sub>C S.cod_char\<^sub>S\<^sub>b\<^sub>C \<omega>_in_hhom
by (intro S.in_homI, auto)
qed
interpret M: monoidal_category S.comp \<open>\<lambda>\<mu>\<nu>. S.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> S.\<alpha> \<open>\<i>[a]\<close>
using S.is_monoidal_category by simp
interpret M\<^sub>\<omega>: monoidal_category S.comp \<open>\<lambda>\<mu>\<nu>. S.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> S.\<alpha> \<omega>
using S\<^sub>\<omega>.is_monoidal_category by simp
interpret M: monoidal_category_with_alternate_unit
S.comp \<open>\<lambda>\<mu>\<nu>. S.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> S.\<alpha> \<open>\<i>[a]\<close> \<omega> ..
have 1: "M\<^sub>\<omega>.unity = w"
- using assms M\<^sub>\<omega>.unity_def S.cod_char S.arr_char
+ using assms M\<^sub>\<omega>.unity_def S.cod_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) S.in_homE S\<^sub>\<omega>.\<omega>_in_vhom)
have 2: "M.unity = a"
- using assms M.unity_def S.cod_char S.arr_char by simp
+ using assms M.unity_def S.cod_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C by simp
have "\<exists>!\<psi>. S.in_hom \<psi> a w \<and> S.iso \<psi> \<and> S.comp \<psi> \<i>[a] = S.comp \<omega> (M.tensor \<psi> \<psi>)"
- using assms 1 2 M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\<omega>.unity_def S.cod_char
+ using assms 1 2 M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\<omega>.unity_def S.cod_char\<^sub>S\<^sub>b\<^sub>C
by simp
show "\<exists>!\<psi>. \<guillemotleft>\<psi> : a \<Rightarrow> w\<guillemotright> \<and> iso \<psi> \<and> \<psi> \<cdot> \<i>[a] = \<omega> \<cdot> (\<psi> \<star> \<psi>)"
proof -
have 1: "\<And>\<psi>. S.in_hom \<psi> a w \<longleftrightarrow> \<guillemotleft>\<psi> : a \<Rightarrow> w\<guillemotright>"
- using assms S.in_hom_char S.arr_char
+ using assms S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) S.ideD(1) S.w_simps(1) S\<^sub>\<omega>.w_simps(1) in_homE
src_dom trg_dom)
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> S.iso \<psi> \<longleftrightarrow> iso \<psi>"
- using assms S.in_hom_char S.arr_char S.iso_char by auto
+ using assms S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C S.iso_char\<^sub>S\<^sub>b\<^sub>C by auto
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> M.tensor \<psi> \<psi> = \<psi> \<star> \<psi>"
- using assms S.in_hom_char S.arr_char S.hcomp_def by simp
+ using assms S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C S.hcomp_def by simp
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> S.comp \<psi> \<i>[a] = \<psi> \<cdot> \<i>[a]"
- using assms S.in_hom_char S.comp_char by auto
+ using assms S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.comp_char by auto
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> S.comp \<omega> (M.tensor \<psi> \<psi>) = \<omega> \<cdot> (\<psi> \<star> \<psi>)"
- using assms S.in_hom_char S.arr_char S.hcomp_def S.comp_char S.dom_char S.cod_char
+ using assms S.in_hom_char\<^sub>S\<^sub>b\<^sub>C S.arr_char\<^sub>S\<^sub>b\<^sub>C S.hcomp_def S.comp_char S.dom_char\<^sub>S\<^sub>b\<^sub>C S.cod_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) M\<^sub>\<omega>.arr_tensor S\<^sub>\<omega>.\<omega>_simps(1) calculation(3) ext)
ultimately show ?thesis
by (metis (no_types, lifting) M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\<omega>.unity_def
S.\<omega>_in_vhom S.in_homE S\<^sub>\<omega>.\<omega>_in_vhom)
qed
qed
end
diff --git a/thys/Bicategory/Tabulation.thy b/thys/Bicategory/Tabulation.thy
--- a/thys/Bicategory/Tabulation.thy
+++ b/thys/Bicategory/Tabulation.thy
@@ -1,6180 +1,6180 @@
(* Title: Tabulation
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Tabulations"
theory Tabulation
imports CanonicalIsos InternalAdjunction
begin
text \<open>
A ``tabulation'' is a kind of bicategorical limit that associates with a 1-cell \<open>r\<close>
a triple \<open>(f, \<rho>, g)\<close>, where \<open>f\<close> and \<open>g\<close> are 1-cells having a common source,
and \<open>\<rho>\<close> is a $2$-cell from \<open>g\<close> to \<open>r \<cdot> f\<close>, such that a certain biuniversal property
is satisfied.
The notion was introduced in a study of bicategories of spans and relations by
Carboni, Kasangian, and Street \cite{carboni-et-al} (hereinafter, ``CKS''),
who named it after a related,
but different notion previously used by Freyd in his study of the algebra of relations.
One can find motivation for the concept of tabulation by considering the problem of
trying to find some kind of universal way of factoring a 1-cell \<open>r\<close>, up to isomorphism,
as the composition \<open>g \<cdot> f\<^sup>*\<close> of a map \<open>g\<close> and the right adjoint \<open>f\<^sup>*\<close> of a map \<open>f\<close>.
In order to be able to express this as a bicategorical limit, CKS consider,
instead of an isomorphism \<open>\<guillemotleft>\<phi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>\<close>, its transpose
\<open>\<rho> : g \<Rightarrow> r \<star> f\<close> under the adjunction \<open>f \<stileturn> f\<^sup>*\<close>.
\<close>
subsection "Definition of Tabulation"
text \<open>
The following locale sets forth the ``signature'' of the data involved in a tabulation,
and establishes some basic facts.
$$\xymatrix{
& \scriptstyle{{\rm src}~g \;=\;{\rm src}~f} \xtwocell[ddd]{}\omit{^\rho}
\ar[ddl] _{g}
\ar[ddr] ^{f}
\\
\\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll] ^{r}
\\
&
}$$
\<close>
locale tabulation_data =
bicategory +
fixes r :: 'a
and \<rho> :: 'a
and f :: 'a
and g :: 'a
assumes ide_base: "ide r"
and ide_leg0: "ide f"
and tab_in_vhom': "\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>"
begin
lemma base_in_hom [intro]:
shows "\<guillemotleft>r : src r \<rightarrow> trg r\<guillemotright>" and "\<guillemotleft>r : r \<Rightarrow> r\<guillemotright>"
using ide_base by auto
lemma base_simps [simp]:
shows "ide r" and "arr r"
and "dom r = r" and "cod r = r"
using ide_base by auto
lemma tab_in_hom [intro]:
shows "\<guillemotleft>\<rho> : src f \<rightarrow> trg r\<guillemotright>" and "\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>"
using tab_in_vhom' src_dom [of \<rho>] trg_dom [of \<rho>] base_in_hom apply auto
by (metis arrI hcomp_simps(1) hcomp_simps(2) in_hhomI not_arr_null
src.is_extensional src.preserves_hom vconn_implies_hpar(1)
vconn_implies_hpar(2) vconn_implies_hpar(3) vconn_implies_hpar(4))
lemma ide_leg1:
shows "ide g"
using tab_in_hom by auto
lemma leg1_in_hom [intro]:
shows "\<guillemotleft>g : src f \<rightarrow> trg r\<guillemotright>" and "\<guillemotleft>g : g \<Rightarrow> g\<guillemotright>"
using ide_leg1 apply auto
using tab_in_hom ide_dom [of \<rho>]
apply (elim conjE in_homE) by auto
lemma leg1_simps [simp]:
shows "ide g" and "arr g"
and "src g = src f" and "trg g = trg r"
and "dom g = g"and "cod g = g"
using ide_leg1 leg1_in_hom by auto
lemma tab_simps [simp]:
shows "arr \<rho>" and "src \<rho> = src f" and "trg \<rho> = trg r"
and "dom \<rho> = g" and "cod \<rho> = r \<star> f"
using tab_in_hom by auto
lemma leg0_in_hom [intro]:
shows "\<guillemotleft>f : src f \<rightarrow> src r\<guillemotright>" and "\<guillemotleft>f : f \<Rightarrow> f\<guillemotright>"
using ide_leg0 apply auto
using tab_in_hom ide_cod [of \<rho>] hseq_char [of r f]
apply (elim conjE in_homE) by auto
lemma leg0_simps [simp]:
shows "ide f" and "arr f"
and "trg f = src r"
and "dom f = f" and "cod f = f"
using ide_leg0 leg0_in_hom by auto
text \<open>
The following function, which composes \<open>\<rho>\<close> with a 2-cell \<open>\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>\<close> to obtain
a 2-cell \<open>\<guillemotleft>(r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"\<close>,
occurs frequently in the sequel.
\<close>
abbreviation (input) composite_cell
where "composite_cell w \<theta> \<equiv> (r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w)"
lemma composite_cell_in_hom:
assumes "ide w" and "\<guillemotleft>w : src u \<rightarrow> src f\<guillemotright>" and "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
shows "\<guillemotleft>composite_cell w \<theta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<rho> \<star> w : g \<star> w \<Rightarrow> (r \<star> f) \<star> w\<guillemotright>"
using assms tab_in_hom
apply (elim conjE in_hhomE in_homE)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[r, f, w] : (r \<star> f) \<star> w \<Rightarrow> r \<star> f \<star> w\<guillemotright>"
using assms ide_base ide_leg0 tab_in_hom by fastforce
show "\<guillemotleft>r \<star> \<theta> : r \<star> f \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using assms ide_base ide_leg0 tab_in_hom by fastforce
qed
text \<open>
We define some abbreviations for various combinations of conditions that occur in the
hypotheses and conclusions of the tabulation axioms.
\<close>
abbreviation (input) uw\<theta>\<omega>
where "uw\<theta>\<omega> u w \<theta> \<omega> \<equiv> ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
abbreviation (input) uw\<theta>\<omega>\<nu>
where "uw\<theta>\<omega>\<nu> u w \<theta> \<omega> \<nu> \<equiv>
ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> = \<omega>"
abbreviation (input) uw\<theta>w'\<theta>'\<beta>
where "uw\<theta>w'\<theta>'\<beta> u w \<theta> w' \<theta>' \<beta> \<equiv>
ide u \<and> ide w \<and> ide w' \<and>
\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright> \<and>
(r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) = (r \<star> \<theta>') \<cdot> \<a>[r, f, w'] \<cdot> (\<rho> \<star> w') \<cdot> \<beta>"
end
text \<open>
CKS define two notions of tabulation.
The first, which they call simply ``tabulation'', is restricted to triples \<open>(f, \<rho>, g)\<close>
where the ``input leg'' \<open>f\<close> is a map, and assumes only a weak form of the biuniversal
property that only applies to \<open>(u, \<omega>, v)\<close> for which u is a map.
The second notion, which they call ``wide tabulation'', concerns arbitrary \<open>(f, \<rho>, g)\<close>,
and assumes a strong form of the biuniversal property that applies to all \<open>(u, \<omega>, v)\<close>.
On its face, neither notion implies the other: ``tabulation'' has the stronger assumption
that \<open>f\<close> is a map, but requires a weaker biuniversal property, and ``wide tabulation''
omits the assumption on \<open>f\<close>, but requires a stronger biuniversal property.
CKS Proposition 1(c) states that if \<open>(f, \<rho>, g)\<close> is a wide tabulation,
then \<open>f\<close> is automatically a map. This is in fact true, but it took me a long time to
reconstruct the details of the proof.
CKS' definition of ``bicategory of spans'' uses their notion ``tabulation'',
presumably because it is only applied in situations where maps are involved and it is more
desirable to have axioms that involve a weaker biuniversal property rather than a stronger one.
However I am more interested in ``wide tabulation'', as it is in some sense the nicer notion,
and since I have had to establish various kinds of preservation results that I don't want
to repeat for both tabulation and wide tabulation, I am using wide tabulation everywhere,
calling it simply ``tabulation''. The fact that the ``input leg'' of a tabulation must
be a map is an essential ingredient throughout.
I have attempted to follow CKS variable naming conventions as much as possible in this
development to avoid confusion when comparing with their paper, even though these are
sometimes at odds with what I have been using elsewhere in this document.
\<close>
locale tabulation =
tabulation_data +
assumes T1: "\<And>u \<omega>.
\<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
and T2: "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w'; \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>;
composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
text \<open>
$$
\textbf{T1:}\qquad\qquad
\xy/u67pt/
\xymatrix{
& {\scriptstyle{{\rm src}~\omega}}
\xlowertwocell[ddddl]{}_{{\rm dom}~\omega\hspace{20pt}}{^\nu}
\xuppertwocell[ddddr]{}^{u}{^\theta}
\ar@ {.>}[dd]^{w}
\\
\\
& \scriptstyle{{\rm src}~g \;=\;{\rm src}~f} \xtwocell[ddd]{}\omit{^\rho}
\ar[ddl] _{g}
\ar[ddr] ^{f}
\\
\\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll] ^{r}
\\
&
}
\endxy
\;\;=\;\;
\xy/u33pt/
\xymatrix{
& \scriptstyle{{\rm src}~\omega} \xtwocell[ddd]{}\omit{^\omega}
\ar[ddl] _{{\rm dom}~\omega}
\ar[ddr] ^{u}
\\
\\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll] ^{r}
\\
&
}
\endxy
$$
\<close>
text \<open>
The following definition includes the additional axiom \<open>T0\<close>, which states that
the ``input leg'' \<open>f\<close> is a map.
\<close>
locale tabulation_data_with_T0 =
tabulation_data +
T0: map_in_bicategory V H \<a> \<i> src trg f
begin
abbreviation \<eta> where "\<eta> \<equiv> T0.\<eta>"
abbreviation \<epsilon> where "\<epsilon> \<equiv> T0.\<epsilon>"
text \<open>
If \<open>\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>\<close> is a 2-cell and \<open>f\<close> is a map, then \<open>\<guillemotleft>T0.trnr\<^sub>\<epsilon> r \<rho> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>\<close>,
where \<open>T0.trnr\<^sub>\<epsilon> r \<rho>\<close> is the adjoint transpose of \<open>\<rho>\<close>.
We will show (CKS Proposition 1(d)) that if \<open>\<rho>\<close> is a tabulation,
then \<open>\<psi> = T0.trnr\<^sub>\<epsilon> r \<rho>\<close> is an isomorphism. However, regardless of whether \<open>\<rho>\<close> is a
tabulation, the mapping \<open>\<rho> \<mapsto> \<psi>\<close> is injective, and we can recover \<open>\<rho>\<close> by the formula:
\<open>\<rho> = (\<psi> \<star> f) \<cdot> T0.trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)\<close>. The proof requires only \<open>T0\<close> and the ``syntactic''
properties of the tabulation data, and in particular does not require the tabulation
conditions \<open>T1\<close> and \<open>T2\<close>. In case \<open>\<rho>\<close> is in fact a tabulation, then this formula can
be interpreted as expressing that \<open>\<rho>\<close> is obtained by transposing the identity
\<open>\<guillemotleft>g \<star> f\<^sup>* : g \<star> f\<^sup>* \<Rightarrow> g \<star> f\<^sup>*\<guillemotright>\<close> to obtain a 2-cell \<open>\<guillemotleft>T0.trnr\<^sub>\<eta> g (g \<star> f\<^sup>*) : g \<Rightarrow> (g \<star> f\<^sup>*) \<star> f\<guillemotright>\<close>
(which may be regarded as the canonical tabulation of \<open>g \<star> f\<^sup>*\<close>), and then composing
with the isomorphism \<open>\<guillemotleft>\<psi> \<star> f : (g \<star> f\<^sup>*) \<star> f \<Rightarrow> r \<star> f\<guillemotright>\<close> to obtain a tabulation of \<open>r\<close>.
This fact will end up being very important in establishing the characterization of
bicategories of spans. Strangely, CKS doesn't make any explicit mention of it.
\<close>
lemma rep_in_hom [intro]:
shows "\<guillemotleft>T0.trnr\<^sub>\<epsilon> r \<rho> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>"
proof (unfold T0.trnr\<^sub>\<epsilon>_def, intro comp_in_homI)
show "\<guillemotleft>\<rho> \<star> f\<^sup>* : g \<star> f\<^sup>* \<Rightarrow> (r \<star> f) \<star> f\<^sup>*\<guillemotright>"
using tab_in_hom T0.antipar(1) by auto
show "\<guillemotleft>\<a>[r, f, f\<^sup>*] : (r \<star> f) \<star> f\<^sup>* \<Rightarrow> r \<star> f \<star> f\<^sup>*\<guillemotright>"
using T0.antipar(1-2) by auto
show "\<guillemotleft>r \<star> \<epsilon> : r \<star> f \<star> f\<^sup>* \<Rightarrow> r \<star> src r\<guillemotright>"
using T0.antipar by auto
show "\<guillemotleft>\<r>[r] : r \<star> src r \<Rightarrow> r\<guillemotright>"
by auto
qed
lemma \<rho>_in_terms_of_rep:
shows "\<rho> = (T0.trnr\<^sub>\<epsilon> r \<rho> \<star> f) \<cdot> T0.trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)"
proof -
have "(T0.trnr\<^sub>\<epsilon> r \<rho> \<star> f) \<cdot> T0.trnr\<^sub>\<eta> g (g \<star> f\<^sup>*) =
(\<r>[r] \<cdot> composite_cell f\<^sup>* \<epsilon> \<star> f) \<cdot> ((g \<star> f\<^sup>*) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] \<cdot> (g \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
unfolding T0.trnr\<^sub>\<epsilon>_def T0.trnr\<^sub>\<eta>_def by simp
text \<open>
$$
\xy/u67pt/
\xymatrix{
& \scriptstyle{{\rm src}~g \;=\;{\rm src}~f}
\ar[ddl]_{g} \ar[ddr]^{f} \xtwocell[ddd]{}\omit{^\rho}
&
\\
\\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll]^{r}
\\
& &
}
\endxy
\;\;=\;\;
\xy/u133pt/
\xymatrix{
& & \scriptstyle{{\rm src}~g \;=\;{\rm src}~f} \ar[dd]
\xtwocell[dddddddl]{}\omit{^\rho}
\xlowertwocell[ddddll]{}_{g}{^{\hspace{20pt}{\rm r}^{-1}[g]}}
\xuppertwocell[ddddrr]{}^{f}{\omit} & &
\xtwocell[dddddddlll]{}\omit{^\epsilon}
\xtwocell[ddddll]{}\omit{^\eta}
\\
& \\
& & \scriptstyle{{\rm src}~g \;=\;{\rm src}~f} \ar[dd]^{f} \ar[ddll]_{g}
& \\
& & & \\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll]^{r}
& &
\scriptstyle{{\rm src}~r} \ar[ll] \ar[uull]_{f^\ast}
\xuppertwocell[llll]{}^{r}<20>{^{\hspace{20pt}{\rm r}[r]}}
\\
& & \\
& & \\
& & & & \\
}
\endxy
$$
\<close>
also have "... = (\<r>[r] \<cdot> composite_cell f\<^sup>* \<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] \<cdot> (g \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
proof -
have "((g \<star> f\<^sup>*) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] = \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f]"
using comp_cod_arr T0.antipar by simp
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> (composite_cell f\<^sup>* \<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] \<cdot> (g \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
using comp_assoc T0.antipar whisker_right [of "f" "\<r>[r]" "composite_cell f\<^sup>* \<epsilon>"]
by fastforce
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<star> f) \<cdot> ((\<rho> \<star> f\<^sup>*) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] \<cdot>
(g \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
using T0.antipar whisker_right [of "f" "(r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*]" "\<rho> \<star> f\<^sup>*"] comp_assoc
by fastforce
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[r, f, f\<^sup>*] \<star> f) \<cdot>
((\<rho> \<star> f\<^sup>*) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] \<cdot> (g \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
using T0.antipar whisker_right [of "f" "r \<star> \<epsilon>" "\<a>[r, f, f\<^sup>*]"] comp_assoc by fastforce
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[r, f, f\<^sup>*] \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] \<cdot> (\<rho> \<star> f\<^sup>* \<star> f) \<cdot> (g \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
proof -
have "((\<rho> \<star> f\<^sup>*) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sup>*, f] = \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] \<cdot> (\<rho> \<star> f\<^sup>* \<star> f)"
using assoc'_naturality [of \<rho> "f\<^sup>*" "f"] T0.antipar by simp
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot>
(\<a>[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] \<cdot>
((r \<star> f) \<star> \<eta>) \<cdot> (\<rho> \<star> src (f)) \<cdot> \<r>\<^sup>-\<^sup>1[g]"
proof -
have "(\<rho> \<star> f\<^sup>* \<star> f) \<cdot> (g \<star> \<eta>) = ((r \<star> f) \<star> \<eta>) \<cdot> (\<rho> \<star> src (f))"
using comp_arr_dom comp_cod_arr T0.antipar interchange [of \<rho> "g" "f\<^sup>* \<star> f" \<eta>]
interchange [of "r \<star> f" \<rho> \<eta> "src (f)"]
by auto
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] \<cdot>
((r \<star> f) \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f] \<cdot> \<rho>"
using runit'_naturality [of \<rho>] by simp
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f] \<cdot>
((r \<star> f) \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f] \<cdot> \<rho>"
proof -
have "(\<a>[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] =
\<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f]"
proof -
have "\<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] =
(\<a>\<^sup>-\<^sup>1[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f]"
using pentagon' [of r "f" "f\<^sup>*" "f"] T0.antipar iso_assoc comp_assoc
invert_side_of_triangle(2)
[of "((\<a>\<^sup>-\<^sup>1[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f])"
"\<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f]" "\<a>\<^sup>-\<^sup>1[r, f, f\<^sup>* \<star> f]"]
by fastforce
hence "(\<a>[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sup>*, f] =
((\<a>[r, f, f\<^sup>*] \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[r, f, f\<^sup>*] \<star> f)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f]"
using comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f]"
proof -
have "(\<a>[r, f, f\<^sup>*] \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[r, f, f\<^sup>*] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] =
((r \<star> f \<star> f\<^sup>*) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f]"
using comp_cod_arr comp_assoc iso_assoc comp_arr_inv T0.antipar
whisker_right [of "f" "\<a>[r, f, f\<^sup>*]" "\<a>\<^sup>-\<^sup>1[r, f, f\<^sup>*]"] comp_assoc_assoc'
by simp
also have "... = \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f]"
using comp_cod_arr T0.antipar by auto
finally show ?thesis
using comp_assoc by metis
qed
finally show ?thesis by blast
qed
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> \<epsilon> \<star> f) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> (r \<star> f \<star> \<eta>) \<cdot> \<a>[r, f, src (f)] \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f] \<cdot> \<rho>"
proof -
have "((r \<star> \<epsilon>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sup>*, f] = \<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> \<epsilon> \<star> f)"
using assoc'_naturality [of r \<epsilon> "f"] by auto
moreover have "\<a>[r, f, f\<^sup>* \<star> f] \<cdot> ((r \<star> f) \<star> \<eta>) = (r \<star> f \<star> \<eta>) \<cdot> \<a>[r, f, src (f)]"
using assoc_naturality [of r "f" \<eta>] T0.antipar by auto
ultimately show ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> (\<epsilon> \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<cdot> (f \<star> \<eta>)) \<cdot> \<a>[r, f, src (f)] \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f] \<cdot> \<rho>"
proof -
have "seq \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] (f \<star> \<eta>)"
using T0.antipar by force
moreover have "seq (\<epsilon> \<star> f) (\<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<cdot> (f \<star> \<eta>))"
using T0.antipar by fastforce
ultimately have "(r \<star> \<epsilon> \<star> f) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]) \<cdot> (r \<star> f \<star> \<eta>) =
r \<star> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<cdot> (f \<star> \<eta>)"
using T0.antipar whisker_left [of r "\<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f]" "f \<star> \<eta>"]
whisker_left [of r "\<epsilon> \<star> f" "\<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<cdot> (f \<star> \<eta>)"]
by auto
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot>
\<a>[r, f, src (f)] \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f] \<cdot> \<rho>"
using T0.triangle_left by simp
also have "... = ((\<r>[r] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> \<l>\<^sup>-\<^sup>1[f])) \<cdot>
((r \<star> \<r>[f]) \<cdot> \<a>[r, f, src (f)] \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f]) \<cdot> \<rho>"
using whisker_left [of r "\<l>\<^sup>-\<^sup>1[f]" "\<r>[f]"] comp_assoc by simp
also have "... = ((r \<star> \<l>[f]) \<cdot> (r \<star> \<l>\<^sup>-\<^sup>1[f])) \<cdot> (\<r>[r \<star> f] \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f]) \<cdot> \<rho>"
using triangle' [of r "f"] runit_hcomp [of r "f"] comp_assoc by simp
also have "... = \<rho>"
proof -
have "(r \<star> \<l>[f]) \<cdot> (r \<star> \<l>\<^sup>-\<^sup>1[f]) = r \<star> f"
using iso_lunit comp_arr_inv' whisker_left [of r "\<l>[f]" "\<l>\<^sup>-\<^sup>1[f]"] by simp
moreover have "(\<r>[r \<star> f] \<cdot> \<r>\<^sup>-\<^sup>1[r \<star> f]) = r \<star> f"
using iso_runit inv_is_inverse comp_arr_inv' by auto
ultimately show ?thesis
using comp_cod_arr by simp
qed
finally show ?thesis by simp
qed
end
text \<open>
The following corresponds to what CKS call ``tabulation''; it supposes axiom \<open>T0\<close>,
but involves weaker versions of \<open>T1\<close> and \<open>T2\<close>. I am calling it ``narrow tabulation''.
\<close>
locale narrow_tabulation =
tabulation_data_with_T0 +
assumes T1: "\<And>u \<omega>.
\<lbrakk> is_left_adjoint u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
and T2: "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> is_left_adjoint u; ide w; ide w';
\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>;
composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
text \<open>
The next few locales are used to bundle up some routine consequences of
the situations described by the hypotheses and conclusions of the tabulation axioms,
so we don't have to keep deriving them over and over again in each context,
and also so as to keep the simplification rules oriented consistently with each other.
\<close>
locale uw\<theta> =
tabulation_data +
fixes u :: 'a
and w :: 'a
and \<theta> :: 'a
assumes uw\<theta>: "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
begin
lemma ide_u:
shows "ide u"
using uw\<theta> by force
lemma u_in_hom [intro]:
shows "\<guillemotleft>u : src u \<rightarrow> src r\<guillemotright>"
using uw\<theta> ide_u ide_cod [of \<theta>] hseq_char [of f w]
apply (intro in_hhomI, simp_all)
by (metis arr_dom in_homE leg0_simps(3) trg_hcomp vconn_implies_hpar(4))
lemma u_simps [simp]:
shows "ide u" and "arr u"
and "trg u = src r"
and "dom u = u" and "cod u = u"
using ide_u u_in_hom by auto
lemma ide_w:
shows "ide w"
using uw\<theta> by auto
lemma w_in_hom [intro]:
shows "\<guillemotleft>w : src u \<rightarrow> src f\<guillemotright>" and "\<guillemotleft>w : w \<Rightarrow> w\<guillemotright>"
proof -
show "\<guillemotleft>w : w \<Rightarrow> w\<guillemotright>"
using ide_w by auto
show "\<guillemotleft>w : src u \<rightarrow> src f\<guillemotright>"
proof
show "arr w" using ide_w by simp
show "src w = src u"
using uw\<theta> ide_dom [of \<theta>] hseq_char [of f w]
by (metis arr_dom in_homE src_cod src_dom hcomp_simps(1))
show "trg w = src f"
using uw\<theta> ide_dom [of \<theta>] hseq_char [of f w]
by (metis arr_dom in_homE)
qed
qed
lemma w_simps [simp]:
shows "ide w" and "arr w"
and "src w = src u" and "trg w = src f"
and "dom w = w" and "cod w = w"
using ide_w w_in_hom by auto
lemma \<theta>_in_hom [intro]:
shows "\<guillemotleft>\<theta> : src u \<rightarrow> src r\<guillemotright>" and "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
proof -
show "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
using uw\<theta> by simp
show "\<guillemotleft>\<theta> : src u \<rightarrow> src r\<guillemotright>"
using uw\<theta> hcomp_simps(1-2)
by (metis arrI in_hhomI u_simps(3) vconn_implies_hpar(1-4))
qed
lemma \<theta>_simps [simp]:
shows "arr \<theta>" and "src \<theta> = src u" and "trg \<theta> = src r"
and "dom \<theta> = f \<star> w" and "cod \<theta> = u"
using \<theta>_in_hom by auto
end
locale uw\<theta>\<omega> =
uw\<theta> +
fixes \<omega> :: 'a
assumes uw\<theta>\<omega>: "uw\<theta>\<omega> u w \<theta> \<omega>"
begin
lemma \<omega>_in_hom [intro]:
shows "\<guillemotleft>\<omega> : src w \<rightarrow> trg r\<guillemotright>" and "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
proof -
show "\<guillemotleft>\<omega> : src w \<rightarrow> trg r\<guillemotright>"
using uw\<theta>\<omega> src_cod [of \<omega>] trg_cod [of \<omega>]
apply (elim conjE in_homE)
by simp
show "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
using uw\<theta>\<omega> by auto
qed
lemma \<omega>_simps [simp]:
shows "arr \<omega>" and "src \<omega> = src w" and "trg \<omega> = trg r"
and "cod \<omega> = r \<star> u"
using \<omega>_in_hom by auto
end
locale uw\<theta>\<omega>\<nu> =
uw\<theta> +
fixes \<omega> :: 'a
and \<nu> :: 'a
assumes uw\<theta>\<omega>\<nu>: "uw\<theta>\<omega>\<nu> u w \<theta> \<omega> \<nu>"
begin
lemma \<nu>_in_hom [intro]:
shows "\<guillemotleft>\<nu> : src u \<rightarrow> trg r\<guillemotright>" and "\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright>"
proof -
show "\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright>"
using uw\<theta>\<omega>\<nu> by auto
show "\<guillemotleft>\<nu> : src u \<rightarrow> trg r\<guillemotright>"
proof
show 1: "arr \<nu>"
using uw\<theta>\<omega>\<nu> by auto
show "src \<nu> = src u"
proof -
have "src (cod \<nu>) = src u"
using uw\<theta>\<omega>\<nu>
by (metis arr_cod hcomp_simps(1) in_homE w_simps(3))
thus ?thesis by simp
qed
show "trg \<nu> = trg r"
proof -
have "trg (cod \<nu>) = trg r"
using uw\<theta>\<omega>\<nu>
by (metis arr_cod hcomp_simps(2) in_homE leg1_simps(4))
thus ?thesis by simp
qed
qed
qed
lemma \<nu>_simps [simp]:
shows "iso \<nu>" and "arr \<nu>" and "src \<nu> = src u" and "trg \<nu> = trg r"
and "cod \<nu> = g \<star> w"
using uw\<theta>\<omega>\<nu> \<nu>_in_hom by auto
sublocale uw\<theta>\<omega>
proof (unfold_locales, intro conjI)
show "ide w"
using uw\<theta>\<omega>\<nu> by simp
show "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
using uw\<theta>\<omega>\<nu> by simp
have "\<guillemotleft>(r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> : dom \<nu> \<Rightarrow> r \<star> u\<guillemotright>"
using ide_base ide_leg0 ide_w by fastforce
thus "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
using uw\<theta>\<omega>\<nu> by auto
qed
end
locale uw\<theta>w'\<theta>' =
tabulation_data V H \<a> \<iota> src trg r \<rho> f g +
uw\<theta>: uw\<theta> V H \<a> \<iota> src trg r \<rho> f g u w \<theta> +
uw'\<theta>': uw\<theta> V H \<a> \<iota> src trg r \<rho> f g u w' \<theta>'
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<iota> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and f :: 'a
and g :: 'a
and u :: 'a
and w :: 'a
and \<theta> :: 'a
and w' :: 'a
and \<theta>' :: 'a
locale uw\<theta>w'\<theta>'\<gamma> =
uw\<theta>w'\<theta>' +
fixes \<gamma> :: 'a
assumes \<gamma>_in_vhom: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
and "\<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
begin
lemma \<gamma>_in_hom [intro]:
shows "\<guillemotleft>\<gamma> : src u \<rightarrow> src f\<guillemotright>" and "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
proof -
show "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
using \<gamma>_in_vhom by simp
show "\<guillemotleft>\<gamma> : src u \<rightarrow> src f\<guillemotright>"
proof
show "arr \<gamma>"
using \<gamma>_in_vhom by auto
show "src \<gamma> = src u"
using \<gamma>_in_vhom src_dom [of \<gamma>]
apply (elim in_homE) by simp
show "trg \<gamma> = src f"
using \<gamma>_in_vhom trg_dom [of \<gamma>]
apply (elim in_homE) by simp
qed
qed
lemma \<gamma>_simps [simp]:
shows "arr \<gamma>"
and "src \<gamma> = src u" and "trg \<gamma> = src f"
and "dom \<gamma> = w" and "cod \<gamma> = w'"
using \<gamma>_in_hom by auto
end
locale uw\<theta>w'\<theta>'\<beta> =
uw\<theta>w'\<theta>' +
fixes \<beta> :: 'a
assumes uw\<theta>w'\<theta>'\<beta>: "uw\<theta>w'\<theta>'\<beta> u w \<theta> w' \<theta>' \<beta>"
begin
lemma \<beta>_in_hom [intro]:
shows "\<guillemotleft>\<beta> : src u \<rightarrow> trg r\<guillemotright>" and "\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
proof -
show "\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
using uw\<theta>w'\<theta>'\<beta> by auto
show "\<guillemotleft>\<beta> : src u \<rightarrow> trg r\<guillemotright>"
using uw\<theta>w'\<theta>'\<beta> src_dom [of \<beta>] trg_dom [of \<beta>] hseq_char [of g w]
apply (elim conjE in_homE) by auto
qed
lemma \<beta>_simps [simp]:
shows "arr \<beta>" and "src \<beta> = src u" and "trg \<beta> = trg r"
and "dom \<beta> = g \<star> w" and "cod \<beta> = g \<star> w'"
using \<beta>_in_hom by auto
end
subsection "Tabulations yield Factorizations"
text \<open>
If \<open>(f, \<rho>, g)\<close> is a (wide) tabulation, then \<open>f\<close> is automatically a map;
this is CKS Proposition 1(c).
The proof sketch provided by CKS is only three lines long, and for a long time I
was only able to prove one of the two triangle identities.
Finally, after gaining a lot of experience with the definitions I saw how to prove
the other.
CKS say nothing about the extra step that seems to be required.
\<close>
context tabulation
begin
text \<open>
The following is used in order to allow us to apply the coherence theorem
to shortcut proofs of equations between canonical arrows.
\<close>
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
lemma satisfies_T0:
shows "is_left_adjoint f"
proof -
text \<open>
The difficulty is filling in details left out by CKS, and accounting for the
fact that they have suppressed unitors and associators everywhere.
In addition, their typography generally uses only parentheses, with no explicit
operation symbols to distinguish between horizontal and vertical composition.
In some cases, for example the statement of T2 in the definition of tabulation,
this makes it difficult for someone not very experienced with the definitions to
reconstruct the correct formulas.
\<close>
text \<open>
CKS say to first apply \<open>T1\<close> with \<open>u = src r\<close>, \<open>v = r\<close>, and \<open>\<rho>' = r\<close>.
However, \<open>\<guillemotleft>r : r \<Rightarrow> r\<guillemotright>\<close>, not \<open>\<guillemotleft>r : r \<Rightarrow> r \<star> src r\<guillemotright>\<close>, so we have to take \<open>\<rho>' = \<r>\<^sup>-\<^sup>1[r]\<close>.
\<close>
obtain f\<^sub>a \<epsilon> \<nu>
where f\<^sub>a: "ide f\<^sub>a \<and> \<guillemotleft>\<epsilon> : f \<star> f\<^sub>a \<Rightarrow> src r\<guillemotright> \<and> \<guillemotleft>\<nu> : r \<Rightarrow> g \<star> f\<^sub>a\<guillemotright> \<and> iso \<nu> \<and>
composite_cell f\<^sub>a \<epsilon> \<cdot> \<nu> = \<r>\<^sup>-\<^sup>1[r]"
using T1 [of "src r" "\<r>\<^sup>-\<^sup>1[r]"] runit'_in_hom [of r] ide_base comp_assoc by auto
have f\<^sub>a': "composite_cell f\<^sub>a \<epsilon> \<cdot> \<nu> = \<r>\<^sup>-\<^sup>1[r]"
using f\<^sub>a by simp
have f\<^sub>a: "ide f\<^sub>a \<and> \<guillemotleft>\<epsilon> : f \<star> f\<^sub>a \<Rightarrow> src r\<guillemotright> \<and> \<guillemotleft>\<nu> : r \<Rightarrow> g \<star> f\<^sub>a\<guillemotright> \<and> iso \<nu>"
using f\<^sub>a by simp
have 1: "src f\<^sub>a = trg f"
using f\<^sub>a f\<^sub>a' comp_assoc
by (metis ide_base leg0_simps(3) runit'_simps(1) seqE src_hcomp vconn_implies_hpar(1)
vseq_implies_hpar(1))
have 2: "trg f\<^sub>a = src g"
using f\<^sub>a by force
have \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star> f\<^sub>a \<Rightarrow> trg f\<guillemotright> \<and> \<guillemotleft>\<epsilon> : trg f \<rightarrow> trg f\<guillemotright> \<and>
arr \<epsilon> \<and> src \<epsilon> = trg f \<and> trg \<epsilon> = trg f \<and> dom \<epsilon> = f \<star> f\<^sub>a \<and> cod \<epsilon> = trg f"
using f\<^sub>a 1 2
by (metis in_hhomI in_homE leg0_simps(3) src_src trg_src vconn_implies_hpar(1-4))
have \<nu>: "\<guillemotleft>\<nu> : r \<Rightarrow> g \<star> f\<^sub>a\<guillemotright> \<and> \<guillemotleft>\<nu> : trg f \<rightarrow> trg g\<guillemotright> \<and>
arr \<nu> \<and> src \<nu> = trg f \<and> trg \<nu> = trg g \<and> dom \<nu> = r \<and> cod \<nu> = g \<star> f\<^sub>a"
using f\<^sub>a by force
text \<open>
Next, CKS say to apply \<open>T2\<close> with \<open>w = trg f\<^sub>a = src f\<close>, \<open>w' = f\<^sub>a \<star> f\<close>, \<open>u = f\<close>,
to obtain the unit and the adjunction conditions, but they don't say explicitly
what to use for \<open>\<theta>\<close>, \<open>\<theta>'\<close>, and \<open>\<beta>\<close>.
We need \<open>\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>\<close> and \<open>\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>\<close>;
\emph{i.e.}~\<open>\<guillemotleft>\<theta> : f \<star> trg f\<^sub>a \<Rightarrow> f\<guillemotright>\<close> and \<open>\<guillemotleft>\<theta>' : f \<star> f\<^sub>a \<star> f \<Rightarrow> f\<guillemotright>\<close>.
Evidently, we may take \<open>\<theta> = \<rho>[f]\<close> and \<open>\<theta>' = \<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]\<close>.
What should be taken for \<open>\<beta>\<close>? Reconstructing this is a little bit more difficult.
\<open>T2\<close> requires \<open>\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>\<close>, hence \<open>\<guillemotleft>\<beta> : g \<star> trg f\<^sub>a \<Rightarrow> g \<star> f\<^sub>a \<star> f\<guillemotright>\<close>.
We have the isomorphism \<open>\<guillemotleft>\<nu> : r \<Rightarrow> g \<star> f\<^sub>a\<guillemotright>\<close> from \<open>T1\<close>. Also \<open>\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>\<close>.
So \<open>\<guillemotleft>\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g] : g \<star> trg f\<^sub>a \<Rightarrow> g \<star> f\<^sub>a \<star> f\<guillemotright>\<close>,
suggesting that we take \<open>\<beta> = \<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g]\<close>.
Now, to apply \<open>T2\<close> we need to satisfy the equation:
\[
\<open>(r \<star> \<theta>) \<cdot> \<a>[r, f, trg f\<^sub>a] \<cdot> (\<rho> \<star> trg f\<^sub>a ) =
(r \<star> \<theta>') \<cdot> \<a>[r, f, f\<^sub>a \<star> f] \<cdot> (\<rho> \<star> f\<^sub>a \<star> f) \<cdot> \<beta>\<close>;
\]
that is, with our choice of \<open>\<theta>\<close>, \<open>\<theta>'\<close>, and \<open>\<beta>\<close>:
\<open>(r \<star> \<r>[f]) \<cdot> \<a>[r, f, trg f\<^sub>a] \<cdot> (\<rho> \<star> trg f\<^sub>a ) =
(r \<star> \<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]) \<cdot> \<a>[r, f, f\<^sub>a \<star> f] \<cdot> (\<rho> \<cdot> (f\<^sub>a \<star> f)) \<cdot>
\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g]\<close>.
It is not too difficult to get the idea of showing that the left-hand side
is equal to \<open>\<rho> \<cdot> \<r>[g]\<close> (note that \<open>trg f\<^sub>a = src f = src g]\<close> and \<open>trg f = src r\<close>),
so we should also try to prove that the right-hand side is equal to this as well.
What we have to work with is the equation:
\[
\<open>\<r>\<^sup>-\<^sup>1[r] = (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sub>a] \<cdot> (\<rho> \<star> f\<^sub>a ) \<cdot> \<nu>\<close>.
\]
After some pondering, I realized that to apply this to the right-hand side of the
equation to be shown requires that we re-associate everything to the left,
so that f stands alone on the right.
\<close>
let ?\<beta> = "\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g]"
let ?\<theta> = "\<r>[f]"
let ?\<theta>' = "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]"
have \<beta>: "\<guillemotleft>?\<beta> : g \<star> src g \<Rightarrow> g \<star> f\<^sub>a \<star> f\<guillemotright> \<and> \<guillemotleft>?\<beta> : src f \<rightarrow> trg g\<guillemotright> \<and>
src ?\<beta> = src g \<and> trg ?\<beta> = trg g \<and> dom ?\<beta> = g \<star> src g \<and> cod ?\<beta> = g \<star> f\<^sub>a \<star> f"
proof -
have 3: "\<guillemotleft>?\<beta> : g \<star> src g \<Rightarrow> g \<star> f\<^sub>a \<star> f\<guillemotright>"
using f\<^sub>a 1 2 by fastforce
moreover have "\<guillemotleft>?\<beta> : src f \<rightarrow> trg g\<guillemotright>"
using 1 2 3 f\<^sub>a by auto
ultimately show ?thesis
by (auto simp add: in_hhom_def)
qed
have \<theta>': "\<guillemotleft>?\<theta>' : f \<star> f\<^sub>a \<star> f \<Rightarrow> f\<guillemotright>"
using f\<^sub>a 1 2 \<epsilon> by fastforce
have A: "composite_cell (trg f\<^sub>a) \<r>[f] = composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> ?\<beta>"
proof -
have "composite_cell (trg f\<^sub>a) \<r>[f] = \<rho> \<cdot> \<r>[g]"
using 2 runit_hcomp runit_naturality [of \<rho>] comp_assoc by simp
also have "... = composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> ?\<beta>"
proof -
have "composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> ?\<beta> =
(composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> \<a>[g, f\<^sub>a, f]) \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g]"
using comp_assoc by simp
also have "... = \<rho> \<cdot> \<r>[g]"
proof -
have "(composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> \<a>[g, f\<^sub>a, f]) \<cdot> (\<nu> \<star> f) = r \<star> f"
proof -
have "(composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> \<a>[g, f\<^sub>a, f]) \<cdot> (\<nu> \<star> f) =
\<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sub>a] \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> \<nu> \<star> f"
proof -
have "(composite_cell (f\<^sub>a \<star> f) ?\<theta>' \<cdot> \<a>[g, f\<^sub>a, f]) \<cdot> (\<nu> \<star> f) =
(r \<star> \<l>[f]) \<cdot> (r \<star> \<epsilon> \<star> f) \<cdot>
composite_cell (f\<^sub>a \<star> f) \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f))"
using f\<^sub>a 1 2 \<epsilon> whisker_left comp_assoc by auto
also have "... = (\<r>[r] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> \<epsilon> \<star> f) \<cdot>
composite_cell (f\<^sub>a \<star> f) \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f))"
using f\<^sub>a 1 2 comp_assoc by (simp add: triangle')
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sub>a, f] \<cdot>
composite_cell (f\<^sub>a \<star> f) \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f))"
proof -
have "\<a>\<^sup>-\<^sup>1[r, src r, f] \<cdot> (r \<star> \<epsilon> \<star> f) = ((r \<star> \<epsilon>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sub>a, f]"
using f\<^sub>a \<epsilon> assoc'_naturality [of r \<epsilon> f] by auto
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot>
(\<a>[r, f, f\<^sub>a] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sub>a, f] \<cdot> (\<rho> \<star> f\<^sub>a \<star> f) \<cdot>
\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f)"
proof -
have "(\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sub>a, f] \<cdot>
composite_cell (f\<^sub>a \<star> f) \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f)) =
(\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot>
(\<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sub>a, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]) \<cdot> \<a>[r, f, f\<^sub>a \<star> f]) \<cdot>
(\<rho> \<star> f\<^sub>a \<star> f) \<cdot> \<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f)"
by (simp add: comp_assoc)
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot>
((\<a>[r, f, f\<^sub>a] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sub>a, f]) \<cdot>
(\<rho> \<star> f\<^sub>a \<star> f) \<cdot> \<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f)"
proof -
have "\<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sub>a, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]) \<cdot> \<a>[r, f, f\<^sub>a \<star> f] =
(\<a>[r, f, f\<^sub>a] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sub>a, f]"
proof -
(* No need to calculate manually, apply the coherence theorem. *)
have "\<a>\<^sup>-\<^sup>1[r, f \<star> f\<^sub>a, f] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]) \<cdot> \<a>[r, f, f\<^sub>a \<star> f] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^sub>a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sub>a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sub>a\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>\<^bold>]\<rbrace>"
using f\<^sub>a 1 2 \<a>'_def \<alpha>_def assoc'_eq_inv_assoc by auto
also have "... = \<lbrace>(\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sub>a\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sub>a\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>\<^bold>]\<rbrace>"
using f\<^sub>a 1 2 by (intro E.eval_eqI, auto)
also have "... = (\<a>[r, f, f\<^sub>a] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sub>a, f]"
using f\<^sub>a 1 2 \<a>'_def \<alpha>_def assoc'_eq_inv_assoc by auto
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[r, f, f\<^sub>a] \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sub>a, f] \<cdot> (\<rho> \<star> f\<^sub>a \<star> f) \<cdot> \<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f)"
by (simp add: comp_assoc)
finally show ?thesis by blast
qed
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot>
(\<a>[r, f, f\<^sub>a] \<star> f) \<cdot> ((\<rho> \<star> f\<^sub>a) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sub>a, f] \<cdot>
\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f)"
proof -
have "\<a>\<^sup>-\<^sup>1[r \<star> f, f\<^sub>a, f] \<cdot> (\<rho> \<star> f\<^sub>a \<star> f) = ((\<rho> \<star> f\<^sub>a) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sub>a, f]"
using f\<^sub>a 1 2 assoc'_naturality [of \<rho> f\<^sub>a f] by auto
thus ?thesis
by (metis comp_assoc)
qed
also have "... = (\<r>[r] \<star> f) \<cdot> ((r \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[r, f, f\<^sub>a] \<star> f) \<cdot>
((\<rho> \<star> f\<^sub>a) \<star> f) \<cdot> (\<nu> \<star> f)"
proof -
have "\<a>\<^sup>-\<^sup>1[g, f\<^sub>a, f] \<cdot> \<a>[g, f\<^sub>a, f] = (g \<star> f\<^sub>a) \<star> f"
using f\<^sub>a 1 2 comp_assoc_assoc' by auto
moreover have "((g \<star> f\<^sub>a) \<star> f) \<cdot> (\<nu> \<star> f) = \<nu> \<star> f"
by (simp add: \<nu> comp_cod_arr)
ultimately show ?thesis
using comp_assoc by metis
qed
also have "... = (\<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sub>a] \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> \<nu>) \<star> f"
proof -
have "arr (\<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sub>a] \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> \<nu>)"
using f\<^sub>a' comp_assoc by auto
thus ?thesis
using whisker_right by fastforce
qed
finally show ?thesis by blast
qed
also have "... = (\<r>[r] \<cdot> \<r>\<^sup>-\<^sup>1[r]) \<star> f"
using f\<^sub>a' comp_assoc by simp
also have "... = r \<star> f"
using ide_base by (simp add: comp_arr_inv')
finally show ?thesis by blast
qed
thus ?thesis
using ide_leg0 ide_leg1 tab_in_hom comp_cod_arr comp_assoc tab_simps(5) arrI
by metis
qed
finally show ?thesis by argo
qed
finally show ?thesis by argo
qed
obtain \<eta> where \<eta>: "\<guillemotleft>\<eta> : trg f\<^sub>a \<Rightarrow> f\<^sub>a \<star> f\<guillemotright> \<and> ?\<beta> = g \<star> \<eta> \<and>
(\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]) \<cdot> (f \<star> \<eta>) = \<r>[f]"
using \<beta> \<theta>' A 1 2 f\<^sub>a runit_in_hom ide_leg0 ide_hcomp src.preserves_ide
T2 [of "trg f\<^sub>a" "f\<^sub>a \<star> f" "\<r>[f]" f "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]" ?\<beta>] comp_assoc
leg1_simps(3)
by metis
have \<eta>': "?\<beta> = g \<star> \<eta> \<and> (\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f]) \<cdot> (f \<star> \<eta>) = \<r>[f]"
using \<eta> by simp
have \<eta>: "\<guillemotleft>\<eta> : trg f\<^sub>a \<Rightarrow> f\<^sub>a \<star> f\<guillemotright> \<and> \<guillemotleft>\<eta> : src f \<rightarrow> src f\<guillemotright> \<and>
arr \<eta> \<and> src \<eta> = src f \<and> trg \<eta> = src f \<and> dom \<eta> = trg f\<^sub>a \<and> cod \<eta> = f\<^sub>a \<star> f"
using \<eta> \<beta> 2 by force
have "adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sub>a \<eta> \<epsilon>"
proof
show "ide f" using ide_leg0 by simp
show "ide f\<^sub>a" using f\<^sub>a by blast
show \<eta>_in_hom: "\<guillemotleft>\<eta> : src f \<Rightarrow> f\<^sub>a \<star> f\<guillemotright>"
using \<eta> 2 by simp
show \<epsilon>_in_hom: "\<guillemotleft>\<epsilon> : f \<star> f\<^sub>a \<Rightarrow> src f\<^sub>a\<guillemotright>"
using f\<^sub>a 1 by simp
show *: "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
using ide_leg0 iso_lunit invert_side_of_triangle(1) \<eta>' comp_assoc by auto
text \<open>
We have proved one of the triangle identities; now we have to show the other.
This part, not mentioned by CKS, took me a while to discover.
Apply \<open>T2\<close> again, this time with the following:
\[\begin{array}{l}
\<open>w = src f \<star> f\<^sub>a\<close>,\\
\<open>\<theta> = (\<epsilon> \<star> \<epsilon>) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a)\<close>,\\
\<open>w' = f\<^sub>a \<star> trg\<close>,\\
\<open>\<theta>' = \<epsilon> \<star> trg f\<close>,\\
\<open>\<beta> = g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]\<close>
\end{array}\]
Then the conditions for \<open>\<gamma>\<close> are satisfied by both
\<open>\<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]\<close> and \<open>(f\<^sub>a \<star> \<epsilon>) \<cdot> \<a>[f\<^sub>a, f, f\<^sub>a] \<cdot> (\<eta> \<star> f\<^sub>a)\<close> so they are equal,
as required.
\<close>
show "(f\<^sub>a \<star> \<epsilon>) \<cdot> \<a>[f\<^sub>a, f, f\<^sub>a] \<cdot> (\<eta> \<star> f\<^sub>a) = \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]"
proof -
let ?u = "trg f \<star> trg f"
let ?w = "src f \<star> f\<^sub>a"
let ?w' = "f\<^sub>a \<star> trg f"
let ?\<theta> = "(\<epsilon> \<star> \<epsilon>) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a)"
let ?\<theta>' = "(\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, trg f]"
let ?\<beta> = "g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]"
let ?\<gamma> = "\<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]"
let ?\<gamma>' = "(f\<^sub>a \<star> \<epsilon>) \<cdot> \<a>[f\<^sub>a, f, f\<^sub>a] \<cdot> (\<eta> \<star> f\<^sub>a)"
have \<theta>_eq': "?\<theta> = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
proof -
have "?\<theta> = (trg f \<star> \<epsilon>) \<cdot> (\<epsilon> \<star> f \<star> f\<^sub>a) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a])) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a)"
using interchange [of "trg f" \<epsilon> \<epsilon> "f \<star> f\<^sub>a"] comp_arr_dom comp_cod_arr comp_assoc
by (simp add: \<epsilon>)
also have "... = (trg f \<star> \<epsilon>) \<cdot> (\<epsilon> \<star> f \<star> f\<^sub>a) \<cdot>
(\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]) \<cdot>
(f \<star> \<eta> \<star> f\<^sub>a)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) =
\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]"
proof -
have "(\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> ((\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a])) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a]) =
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> ide_leg0 iso_assoc
invert_side_of_triangle(1)
[of "((\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]) \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a])"
"\<a>\<^sup>-\<^sup>1[f \<star> f\<^sub>a, f, f\<^sub>a]" "\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]"]
pentagon' comp_assoc by auto
hence "(\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> ((\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a])) =
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a])"
using 1 2 \<open>ide f\<^sub>a\<close>
invert_side_of_triangle(2)
[of "\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]" "\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> ((\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a])"
"f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a]"]
by auto
thus ?thesis
using comp_assoc by simp
qed
thus ?thesis by simp
qed
also have "... = (trg f \<star> \<epsilon>) \<cdot> ((\<epsilon> \<star> f \<star> f\<^sub>a) \<cdot> \<a>[f \<star> f\<^sub>a, f, f\<^sub>a]) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a] \<cdot> (f \<star> \<eta> \<star> f\<^sub>a)"
using comp_assoc by simp
also have "... = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot>
((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (f \<star> \<eta>) \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
proof -
have "((\<epsilon> \<star> f \<star> f\<^sub>a) \<cdot> \<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a) =
(\<a>[trg f, f, f\<^sub>a] \<cdot> ((\<epsilon> \<star> f) \<star> f\<^sub>a)) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot>
((f \<star> \<eta>) \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using assoc_naturality [of \<epsilon> f f\<^sub>a] assoc'_naturality [of f \<eta> f\<^sub>a]
by (simp add: 2 \<epsilon> \<eta> \<open>ide f\<^sub>a\<close> comp_assoc)
also have "... = \<a>[trg f, f, f\<^sub>a] \<cdot>
(((\<epsilon> \<star> f) \<star> f\<^sub>a) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> ((f \<star> \<eta>) \<star> f\<^sub>a)) \<cdot>
\<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using comp_assoc by simp
also have "... = \<a>[trg f, f, f\<^sub>a] \<cdot>
((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (f \<star> \<eta>) \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using \<eta>' comp_assoc whisker_right \<open>ide f\<^sub>a\<close> comp_null(2) ide_leg0 ext
runit_simps(1)
by metis
finally show ?thesis
using comp_assoc by simp
qed
also have "... = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using * by simp
finally show ?thesis by simp
qed
have \<theta>_eq: "?\<theta> = (\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, src f\<^sub>a] \<cdot> (f \<star> ?\<gamma>)"
proof -
have "?\<theta> = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using \<theta>_eq' by simp
also have "... =
(trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<star> f\<^sub>a) \<cdot> (\<r>[f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using \<open>ide f\<^sub>a\<close> whisker_right comp_assoc by auto
also have "... = (trg f \<star> \<epsilon>) \<cdot> ((\<a>[trg f, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[trg f, f, f\<^sub>a]) \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> f\<^sub>a])) \<cdot>
(f \<star> \<l>[f\<^sub>a])"
using 2 \<open>ide f\<^sub>a\<close> lunit_hcomp [of f f\<^sub>a] invert_side_of_triangle(2) triangle'
comp_assoc
by auto
also have "... = (trg f \<star> \<epsilon>) \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> f\<^sub>a] \<cdot> (f \<star> \<l>[f\<^sub>a])"
using f\<^sub>a 2 comp_cod_arr iso_assoc comp_arr_inv lunit_hcomp(2) lunit_hcomp(4)
ide_leg0 leg1_simps(3)
by metis
also have "... = \<l>\<^sup>-\<^sup>1[trg f] \<cdot> \<epsilon> \<cdot> (f \<star> \<l>[f\<^sub>a])"
using \<epsilon> lunit'_naturality comp_assoc by metis
also have "... = \<r>\<^sup>-\<^sup>1[trg f] \<cdot> \<epsilon> \<cdot> (f \<star> \<l>[f\<^sub>a])"
using unitor_coincidence by simp
also have "... = (\<epsilon> \<star> trg f) \<cdot> \<r>\<^sup>-\<^sup>1[f \<star> f\<^sub>a] \<cdot> (f \<star> \<l>[f\<^sub>a])"
using \<epsilon> runit'_naturality comp_assoc by metis
also have "... = (\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, src f\<^sub>a] \<cdot> (f \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a]) \<cdot> (f \<star> \<l>[f\<^sub>a])"
using 2 \<open>ide f\<^sub>a\<close> runit_hcomp(2) comp_assoc by auto
also have "... = (\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, src f\<^sub>a] \<cdot> (f \<star> ?\<gamma>)"
using 2 \<open>ide f\<^sub>a\<close> whisker_left by simp
finally show ?thesis by simp
qed
have \<theta>: "\<guillemotleft>?\<theta> : f \<star> ?w \<Rightarrow> ?u\<guillemotright>"
using 1 2 \<open>ide f\<^sub>a\<close> \<eta>_in_hom \<epsilon> by fastforce
have \<theta>': "\<guillemotleft>?\<theta>' : f \<star> ?w' \<Rightarrow> ?u\<guillemotright>"
using f\<^sub>a 1 2 \<epsilon> by auto
have ww': "ide ?w \<and> ide ?w'"
by (simp add: 1 2 \<open>ide f\<^sub>a\<close>)
have "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w \<Rightarrow> ?w'\<guillemotright> \<and> ?\<beta> = g \<star> \<gamma> \<and> ?\<theta> = ?\<theta>' \<cdot> (f \<star> \<gamma>)"
proof -
have "\<guillemotleft>?\<beta> : g \<star> ?w \<Rightarrow> g \<star> ?w'\<guillemotright>"
using \<open>ide f\<^sub>a\<close> 1 2 by auto
moreover have "composite_cell ?w ?\<theta> = composite_cell ?w' ?\<theta>' \<cdot> ?\<beta>"
proof -
have "composite_cell ?w' ?\<theta>' \<cdot> ?\<beta> =
composite_cell ?w ((\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, src f\<^sub>a] \<cdot> (f \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]))"
proof -
have "\<a>[r, f, f\<^sub>a \<star> trg f] \<cdot> (\<rho> \<star> f\<^sub>a \<star> trg f) \<cdot> (g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]) =
composite_cell ?w (f \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a])"
proof -
have "\<a>[r, f, f\<^sub>a \<star> trg f] \<cdot> (\<rho> \<star> f\<^sub>a \<star> trg f) \<cdot> (g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]) =
(\<a>[r, f, f\<^sub>a \<star> trg f] \<cdot> ((r \<star> f) \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a])) \<cdot> (\<rho> \<star> src f \<star> f\<^sub>a)"
proof -
have "(\<rho> \<star> f\<^sub>a \<star> trg f) \<cdot> (g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]) = \<rho> \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]"
using interchange [of \<rho> g "f\<^sub>a \<star> trg f" "\<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]"]
comp_arr_dom comp_cod_arr 1 2 \<open>ide f\<^sub>a\<close>
by simp
also have "... = ((r \<star> f) \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]) \<cdot> (\<rho> \<star> src f \<star> f\<^sub>a)"
proof -
have "seq (f\<^sub>a \<star> trg f) (\<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a])"
using f\<^sub>a 1 2 ww' by auto
thus ?thesis
using interchange comp_arr_dom comp_cod_arr 1 2 \<open>ide f\<^sub>a\<close>
by (metis ww' comp_ide_arr dom_comp leg1_simps(3)
lunit_simps(4) tab_simps(1) tab_simps(5))
qed
finally show ?thesis
using comp_assoc by simp
qed
also have "... = composite_cell ?w (f \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a])"
using assoc_naturality [of r f "\<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]"] 1 2 \<open>ide f\<^sub>a\<close> comp_assoc by simp
finally show ?thesis by simp
qed
hence "composite_cell ?w' ?\<theta>' \<cdot> ?\<beta> =
((r \<star> (\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, trg f]) \<cdot> (r \<star> f \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a])) \<cdot>
\<a>[r, f, src f \<star> f\<^sub>a] \<cdot> (\<rho> \<star> src f \<star> f\<^sub>a)"
using comp_assoc by simp
also have
"... = composite_cell ?w (((\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, trg f]) \<cdot> (f \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]))"
using whisker_left 1 2 \<open>ide f\<^sub>a\<close> ide_base
by (metis \<open>\<guillemotleft>(\<epsilon> \<star> \<epsilon>) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a) :
f \<star> src f \<star> f\<^sub>a \<Rightarrow> trg f \<star> trg f\<guillemotright>\<close>
\<theta>_eq arrI comp_assoc)
finally show ?thesis
using comp_assoc by (simp add: "1")
qed
also have "... = composite_cell ?w ?\<theta>"
using \<theta>_eq by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using ww' \<theta> \<theta>' T2 [of ?w ?w' ?\<theta> ?u ?\<theta>' ?\<beta>] comp_assoc by metis
qed
moreover have "\<guillemotleft>?\<gamma> : ?w \<Rightarrow> ?w'\<guillemotright> \<and> ?\<beta> = g \<star> ?\<gamma> \<and> ?\<theta> = ?\<theta>' \<cdot> (f \<star> ?\<gamma>)"
using 1 2 \<open>ide f\<^sub>a\<close> \<theta>_eq comp_assoc by auto
moreover have "\<guillemotleft>?\<gamma>' : ?w \<Rightarrow> ?w'\<guillemotright> \<and> ?\<beta> = g \<star> ?\<gamma>' \<and> ?\<theta> = ?\<theta>' \<cdot> (f \<star> ?\<gamma>')"
proof (intro conjI)
show "\<guillemotleft>?\<gamma>' : ?w \<Rightarrow> ?w'\<guillemotright>"
using 1 2 f\<^sub>a \<eta>_in_hom \<epsilon>_in_hom by fastforce
show "?\<beta> = g \<star> ?\<gamma>'"
text \<open>
This equation is not immediate.
To show it, we have to recall the properties from the construction of \<open>\<epsilon>\<close> and \<open>\<eta>\<close>.
Use the property of \<open>\<eta>\<close> to replace \<open>g \<star> \<eta> \<star> f\<^sub>a\<close> by a 2-cell involving
\<open>\<epsilon>\<close>, \<open>\<rho>\<close>, and \<open>\<nu>\<close>.
Use the property \<open>(r \<star> \<epsilon>) \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> \<nu> = \<r>[r]\<close> from the construction of \<open>\<epsilon>\<close> to
eliminate \<open>\<epsilon>\<close> and \<open>\<rho>\<close> in favor of inv \<open>\<nu>\<close> and canonical isomorphisms.
Cancelling \<open>\<nu>\<close> and inv \<open>\<nu>\<close> leaves the canonical 2-cell \<open>g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a] \<cdot> \<l>[f\<^sub>a]\<close>.
\<close>
proof -
have "g \<star> ?\<gamma>' = (g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot> (g \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (g \<star> \<eta> \<star> f\<^sub>a)"
using 1 2 \<open>ide f\<^sub>a\<close> \<epsilon> \<eta> whisker_left
by (metis \<open>\<guillemotleft>?\<gamma>' : ?w \<Rightarrow> ?w'\<guillemotright>\<close> arrI ide_leg1 seqE)
also have "... = (g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot> (g \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (g \<star> \<eta> \<star> f\<^sub>a) \<cdot>
\<a>[g, src f, f\<^sub>a] \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> \<eta> comp_arr_dom hseq_char comp_assoc_assoc'
by simp
also have "... = (g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot> (g \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> ((g \<star> \<eta> \<star> f\<^sub>a) \<cdot>
\<a>[g, src f, f\<^sub>a]) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using comp_assoc by simp
also have "... = (g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot> (g \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot>
(\<a>[g, f\<^sub>a \<star> f, f\<^sub>a] \<cdot> ((g \<star> \<eta>) \<star> f\<^sub>a)) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> \<epsilon> \<eta> assoc_naturality [of g \<eta> f\<^sub>a] by simp
also have "... = (g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot> (g \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> \<a>[g, f\<^sub>a \<star> f, f\<^sub>a] \<cdot>
(\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using \<eta>' comp_assoc by simp
also have "... = (g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot>
((g \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> \<a>[g, f\<^sub>a \<star> f, f\<^sub>a] \<cdot> (\<a>[g, f\<^sub>a, f] \<star> f\<^sub>a)) \<cdot>
((\<nu> \<star> f) \<star> f\<^sub>a) \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> (\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
proof -
have "\<a>[g, f\<^sub>a, f] \<cdot> (\<nu> \<star> f) \<cdot> \<rho> \<cdot> \<r>[g] \<star> f\<^sub>a =
(\<a>[g, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> ((\<nu> \<star> f) \<star> f\<^sub>a) \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> (\<r>[g] \<star> f\<^sub>a)"
using 1 2 \<open>ide f\<^sub>a\<close> \<beta> \<epsilon> \<eta> whisker_right by (metis arrI seqE)
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((g \<star> f\<^sub>a \<star> \<epsilon>) \<cdot>
\<a>[g, f\<^sub>a, f \<star> f\<^sub>a]) \<cdot> (\<a>[g \<star> f\<^sub>a, f, f\<^sub>a] \<cdot>
((\<nu> \<star> f) \<star> f\<^sub>a)) \<cdot> (\<rho> \<star> f\<^sub>a) \<cdot> (\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> pentagon comp_assoc by simp
also have "... = (\<a>[g, f\<^sub>a, trg f] \<cdot> ((g \<star> f\<^sub>a) \<star> \<epsilon>)) \<cdot>
((\<nu> \<star> f \<star> f\<^sub>a) \<cdot> \<a>[r, f, f\<^sub>a]) \<cdot>
(\<rho> \<star> f\<^sub>a) \<cdot> (\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> assoc_naturality [of g f\<^sub>a \<epsilon>] assoc_naturality [of \<nu> f f\<^sub>a]
by (simp add: \<epsilon> \<nu>)
also have "... = \<a>[g, f\<^sub>a, trg f] \<cdot> (((g \<star> f\<^sub>a) \<star> \<epsilon>) \<cdot> (\<nu> \<star> f \<star> f\<^sub>a)) \<cdot> \<a>[r, f, f\<^sub>a] \<cdot>
(\<rho> \<star> f\<^sub>a) \<cdot> (\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> assoc_naturality [of g f\<^sub>a \<epsilon>] assoc_naturality [of \<nu> f f\<^sub>a]
comp_assoc
by simp
also have "... = \<a>[g, f\<^sub>a, trg f] \<cdot> (\<nu> \<star> trg f) \<cdot>
composite_cell f\<^sub>a \<epsilon> \<cdot>
(\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
proof -
have "((g \<star> f\<^sub>a) \<star> \<epsilon>) \<cdot> (\<nu> \<star> f \<star> f\<^sub>a) = \<nu> \<star> \<epsilon>"
using 1 2 \<open>ide f\<^sub>a\<close> \<nu> \<epsilon> interchange [of "g \<star> f\<^sub>a" \<nu> \<epsilon> "f \<star> f\<^sub>a"]
comp_arr_dom comp_cod_arr
by simp
also have "... = (\<nu> \<star> trg f) \<cdot> (r \<star> \<epsilon>)"
using \<open>ide f\<^sub>a\<close> \<nu> \<epsilon> interchange [of \<nu> r "trg f" \<epsilon>] comp_arr_dom comp_cod_arr
by simp
finally show ?thesis
using comp_assoc by simp
qed
also have "... = \<a>[g, f\<^sub>a, trg f] \<cdot> ((((\<nu> \<star> trg f) \<cdot> \<r>\<^sup>-\<^sup>1[r]) \<cdot> inv \<nu>) \<cdot> (\<r>[g] \<star> f\<^sub>a)) \<cdot>
\<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using ide_base f\<^sub>a' comp_assoc f\<^sub>a runit'_simps(1) invert_side_of_triangle(2)
comp_assoc
by presburger
also have "... = \<a>[g, f\<^sub>a, trg f] \<cdot> \<r>\<^sup>-\<^sup>1[g \<star> f\<^sub>a] \<cdot> (\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
proof -
have "((\<nu> \<star> trg f) \<cdot> \<r>\<^sup>-\<^sup>1[r]) \<cdot> inv \<nu> = \<r>\<^sup>-\<^sup>1[g \<star> f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> \<nu> ide_base runit'_naturality [of \<nu>] comp_arr_dom
by (metis f\<^sub>a ide_compE inv_is_inverse inverse_arrowsE comp_assoc
runit'_simps(1) runit'_simps(4))
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((\<a>[g, f\<^sub>a, trg f] \<cdot> \<a>\<^sup>-\<^sup>1[g, f\<^sub>a, src f\<^sub>a]) \<cdot>
(g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a])) \<cdot> (\<r>[g] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[g, src f, f\<^sub>a]"
using f\<^sub>a "2" runit_hcomp \<open>ide f\<^sub>a\<close> comp_assoc by simp
also have "... = (g \<star> \<r>\<^sup>-\<^sup>1[f\<^sub>a]) \<cdot> (g \<star> \<l>[f\<^sub>a])"
using 1 2 comp_cod_arr \<open>ide f\<^sub>a\<close> comp_assoc_assoc' triangle' by simp
also have "... = ?\<beta>"
using 2 \<open>ide f\<^sub>a\<close> whisker_left by simp
finally show ?thesis by simp
qed
show "?\<theta> = ?\<theta>' \<cdot> (f \<star> ?\<gamma>')"
proof -
have "((\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, trg f]) \<cdot> (f \<star> (f\<^sub>a \<star> \<epsilon>) \<cdot> \<a>[f\<^sub>a, f, f\<^sub>a] \<cdot> (\<eta> \<star> f\<^sub>a)) =
((\<epsilon> \<star> trg f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, trg f]) \<cdot> (f \<star> f\<^sub>a \<star> \<epsilon>) \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a)"
using 1 2 \<open>ide f\<^sub>a\<close> \<epsilon> \<eta> whisker_left
by (metis \<open>\<guillemotleft>(f\<^sub>a \<star> \<epsilon>) \<cdot> \<a>[f\<^sub>a, f, f\<^sub>a] \<cdot> (\<eta> \<star> f\<^sub>a) : src f \<star> f\<^sub>a \<Rightarrow> f\<^sub>a \<star> trg f\<guillemotright>\<close>
arrI ide_leg0 seqE)
also have
"... = (\<epsilon> \<star> trg f) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, trg f] \<cdot> (f \<star> f\<^sub>a \<star> \<epsilon>)) \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot> (f \<star> \<eta> \<star> f\<^sub>a)"
using comp_assoc by simp
also have "... = ((\<epsilon> \<star> trg f) \<cdot> ((f \<star> f\<^sub>a) \<star> \<epsilon>)) \<cdot>
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) \<cdot>
(f \<star> \<eta> \<star> f\<^sub>a)"
using 1 2 \<open>ide f\<^sub>a\<close> \<epsilon> assoc'_naturality [of f f\<^sub>a \<epsilon>] comp_assoc by simp
also have "... = (trg f \<star> \<epsilon>) \<cdot> (\<epsilon> \<star> f \<star> f\<^sub>a) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a])) \<cdot>
(f \<star> \<eta> \<star> f\<^sub>a)"
using 1 2 \<open>ide f\<^sub>a\<close> \<epsilon> interchange [of \<epsilon> "f \<star> f\<^sub>a" "trg f" \<epsilon>]
interchange [of "trg f" \<epsilon> \<epsilon> "f \<star> f\<^sub>a"] comp_arr_dom comp_cod_arr comp_assoc
by simp
also have "... = (trg f \<star> \<epsilon>) \<cdot> ((\<epsilon> \<star> f \<star> f\<^sub>a) \<cdot>
(\<a>[f \<star> f\<^sub>a, f, f\<^sub>a]) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]) \<cdot>
(f \<star> \<eta> \<star> f\<^sub>a))"
proof -
have "\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a] \<cdot> (f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]) =
\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]"
proof -
have A: "(\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a] \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a]) =
\<a>\<^sup>-\<^sup>1[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> pentagon' comp_assoc by fastforce
hence B: "\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a] \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a]) =
\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]"
using A 1 2 \<open>ide f\<^sub>a\<close>
invert_side_of_triangle(1)
[of "(\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a] \<cdot> (f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a])"
"\<a>\<^sup>-\<^sup>1[f \<star> f\<^sub>a, f, f\<^sub>a]" "\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]"]
by auto
show ?thesis
proof -
have C: "iso (f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a])"
using 1 2 \<open>ide f\<^sub>a\<close> by simp
moreover have "inv (f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a]) = f \<star> \<a>[f\<^sub>a, f, f\<^sub>a]"
using C 1 2 \<open>ide f\<^sub>a\<close> by fastforce
ultimately show ?thesis
using B 1 2 \<open>ide f\<^sub>a\<close> comp_assoc
invert_side_of_triangle(2)
[of "\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f \<star> f\<^sub>a]"
"\<a>[f \<star> f\<^sub>a, f, f\<^sub>a] \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a \<star> f, f\<^sub>a]"
"f \<star> \<a>\<^sup>-\<^sup>1[f\<^sub>a, f, f\<^sub>a]"]
by simp
qed
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = (trg f \<star> \<epsilon>) \<cdot> (\<a>[trg f, f, f\<^sub>a] \<cdot>
((\<epsilon> \<star> f) \<star> f\<^sub>a)) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> ((f \<star> \<eta>) \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> \<open>ide f\<close> \<eta> \<epsilon> assoc_naturality [of \<epsilon> f f\<^sub>a]
assoc'_naturality [of f \<eta> f\<^sub>a] comp_assoc
by simp
also have "... = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot>
(((\<epsilon> \<star> f) \<star> f\<^sub>a) \<cdot> (\<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<star> f\<^sub>a) \<cdot> ((f \<star> \<eta>) \<star> f\<^sub>a)) \<cdot>
\<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using comp_assoc by simp
also have "... = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot>
((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sub>a, f] \<cdot> (f \<star> \<eta>) \<star> f\<^sub>a) \<cdot>
\<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using 1 2 \<open>ide f\<^sub>a\<close> \<open>ide f\<close> \<eta> \<epsilon> whisker_right
by (metis (full_types) * \<theta> \<theta>_eq' arrI hseqE seqE)
also have "... = (trg f \<star> \<epsilon>) \<cdot> \<a>[trg f, f, f\<^sub>a] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> f\<^sub>a) \<cdot> \<a>\<^sup>-\<^sup>1[f, src f, f\<^sub>a]"
using * by simp
also have "... = ?\<theta>"
using \<theta>_eq' by simp
finally show ?thesis by simp
qed
qed
ultimately show "?\<gamma>' = ?\<gamma>" by blast
qed
qed
thus ?thesis
using adjoint_pair_def by auto
qed
sublocale tabulation_data_with_T0
using satisfies_T0 by (unfold_locales, simp)
sublocale narrow_tabulation
using adjoint_pair_antipar(1) T1 T2
by (unfold_locales, auto)
end
text \<open>
A tabulation \<open>(f, \<rho>, g)\<close> of \<open>r\<close> yields an isomorphism \<open>\<guillemotleft>\<psi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>\<close>
via adjoint transpose.
The proof requires \<open>T0\<close>, in order to obtain \<open>\<psi>\<close> as the transpose of \<open>\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>\<close>.
However, it uses only the weaker versions of \<open>T1\<close> and \<open>T2\<close>.
\<close>
context narrow_tabulation
begin
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
text \<open>
The following is CKS Proposition 1(d), with the statement refined to incorporate
the canonical isomorphisms that they omit.
Note that we can easily show using \<open>T1\<close> that there is some 1-cell \<open>f\<^sub>a\<close> and isomorphism \<open>\<psi>\<close>
such that \<open>\<guillemotleft>\<psi> : f \<star> f\<^sub>a \<Rightarrow> r\<guillemotright>\<close> (this was already part of the proof that a tabulation
satisfies \<open>T0\<close>). The more difficult content in the present result is that we may
actually take \<open>f\<^sub>a\<close> to be the left adjoint \<open>f\<^sup>*\<close> of \<open>f\<close>.
\<close>
lemma yields_isomorphic_representation:
shows "\<guillemotleft>T0.trnr\<^sub>\<epsilon> r \<rho> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>" and "iso (T0.trnr\<^sub>\<epsilon> r \<rho>)"
proof -
text \<open>
As stated in CKS, the first step of the proof is:
\begin{quotation}
``Apply \<open>T1\<close> with \<open>X = A\<close>, \<open>u = 1\<^sub>A\<close>, \<open>v = r\<close>, \<open>\<omega> = 1\<^sub>R\<close>, to obtain \<open>f'\<close>, \<open>\<theta>': ff' \<Rightarrow> 1\<^sub>A\<close>,
\<open>\<nu> : r \<simeq> g f'\<close> with \<open>1\<^sub>R = (r\<theta>')(\<rho>f')\<nu>\<close>.''
\end{quotation}
In our nomenclature: \<open>X = trg f\<close>, \<open>u = trg f\<close>, \<open>v = r\<close>, but \<open>\<omega> = src f\<close>
does not make any sense, since we need \<open>\<guillemotleft>\<omega> : v \<Rightarrow> r \<star> u\<guillemotright>\<close>. We have to take \<open>\<omega> = \<r>\<^sup>-\<^sup>1[r]\<close>.
It is not clear whether this is a typo, or whether it is a consequence of CKS having
suppressed all canonical isomorphisms (unitors, in this case). The resulting equation
obtained via T1 is:
\[
\<open>\<r>\<^sup>-\<^sup>1[r] = (r \<star> \<theta>') \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu>\<close>,
\]
which has \<open>\<r>\<^sup>-\<^sup>1[r]\<close> on the left-hand side, rather than \<open>1\<^sub>R\<close>, as in CKS.
Also, we have inserted the omitted associativity.
\<close>
obtain w \<theta>' \<nu> where w\<theta>'\<nu>: "ide w \<and> \<guillemotleft>\<theta>' : f \<star> w \<Rightarrow> src r\<guillemotright> \<and> \<guillemotleft>\<nu> : r \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta>' \<cdot> \<nu> = \<r>\<^sup>-\<^sup>1[r]"
using ide_base obj_is_self_adjoint T1 [of "src r" "\<r>\<^sup>-\<^sup>1[r]"] comp_assoc by auto
interpret uw\<theta>\<omega>\<nu> V H \<a> \<i> src trg r \<rho> f g \<open>src r\<close> w \<theta>' \<open>\<r>\<^sup>-\<^sup>1[r]\<close> \<nu>
using ide_base tab_in_hom w\<theta>'\<nu> comp_assoc by (unfold_locales, auto)
text \<open>
CKS now say:
\begin{quotation}
``Apply \<open>T2\<close> with \<open>u = 1\<^sub>A\<close>, \<open>w = f\<^sup>*\<close>, \<open>w' = f'\<close>, \<open>\<theta> = \<epsilon>: ff\<^sup>* \<Rightarrow> 1\<close>, \<open>\<theta>': ff' \<Rightarrow> 1\<close>,
\<open>\<beta> = \<nu>(r\<epsilon>)(\<rho>f\<^sup>*)\<close> to obtain \<open>\<gamma> : f\<^sup>* \<Rightarrow> f'\<close> with \<open>g\<gamma> = \<nu>(r\<epsilon>)(\<rho>f\<^sup>*)\<epsilon> = \<theta>'(f\<gamma>).\<close>''
\end{quotation}
The last equation is mysterious, but upon consideration one eventually realizes
that it is definitely a typo, and what is meant is ``\<open>g\<gamma> = \<nu>(r\<epsilon>)(\<rho>f\<^sup>*)\<close>, \<open>\<epsilon> = \<theta>'(f\<gamma>)\<close>''.
So, we take \<open>u = trg f\<close>, \<open>w = f\<^sup>*\<close>, \<open>w' = w\<close>, \<open>\<theta>'\<close> as obtained from \<open>T1\<close>, \<open>\<theta> = \<epsilon>\<close>,
and \<open>\<beta> = \<nu> \<cdot> \<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot> (\<rho> \<star> f\<^sup>*)\<close>.
(CKS mention neither the unitor term \<open>\<r>[r]\<close> nor the associativity \<open>\<a>[r, f, f\<^sup>*]\<close>
which are required for the expression for \<open>\<beta>\<close> to make sense.)
\<close>
let ?\<psi> = "\<r>[r] \<cdot> composite_cell f\<^sup>* \<epsilon>"
show \<psi>_in_hom: "\<guillemotleft>T0.trnr\<^sub>\<epsilon> r \<rho> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>"
using ide_base T0.trnr\<^sub>\<epsilon>_def rep_in_hom by simp
have A: "\<guillemotleft>\<nu> \<cdot> ?\<psi> : g \<star> f\<^sup>* \<Rightarrow> g \<star> w\<guillemotright>"
using ide_base T0.antipar hseq_char T0.trnr\<^sub>\<epsilon>_def rep_in_hom w\<theta>'\<nu>
apply (intro comp_in_homI') by auto
have B: "composite_cell f\<^sup>* \<epsilon> = composite_cell w \<theta>' \<cdot> \<nu> \<cdot> ?\<psi>"
using ide_base T0.antipar w\<theta>'\<nu> comp_assoc
by (metis A arrI invert_side_of_triangle(1) iso_runit)
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : f\<^sup>* \<Rightarrow> w\<guillemotright> \<and> \<nu> \<cdot> ?\<psi> = g \<star> \<gamma> \<and> \<epsilon> = \<theta>' \<cdot> (f \<star> \<gamma>)"
using A B T0.counit_in_hom obj_is_self_adjoint T0.antipar comp_assoc
T2 [of "trg f" "f\<^sup>*" w \<epsilon> \<theta>' "\<nu> \<cdot> \<r>[r] \<cdot> composite_cell f\<^sup>* \<epsilon>"]
by auto
have trg_\<gamma>_eq: "trg \<gamma> = trg w"
using \<gamma> by fastforce
text \<open>
CKS say:
\begin{quotation}
``The last equation implies \<open>\<gamma>: f\<^sup>* \<Rightarrow> f'\<close> is a split monic (coretraction), while
the calculation:
\begin{eqnarray*}
\<open>(g\<gamma>)(gf\<^sup>*\<theta>')(g\<eta>f')\<close> &\<open>=\<close>& \<open>\<nu>(r\<epsilon>)(\<rho>f\<^sup>*)(gf\<^sup>*\<theta>')(g\<eta>f')\<close>\\
&\<open>=\<close>& \<open>\<nu>(r\<epsilon>)(rff\<^sup>*\<theta>')(\<rho>f\<^sup>*ff')(g\<eta>f')\<close>\\
&\<open>=\<close>& \<open>\<nu>(r\<theta>')(r\<epsilon>ff')(rf\<eta>f')(\<rho>f')\<close>\\
&\<open>=\<close>& \<open>\<nu>(r\<theta>')(\<rho>f') = 1\<^sub>g\<^sub>f\<^sub>'\<close>,
\end{eqnarray*}
shows that \<open>g\<gamma>\<close> is a split epic. So \<open>g\<gamma> = \<nu>(r\<epsilon>)(\<rho>f\<^sup>*): gf\<^sup>* \<Rightarrow> gf'\<close> is invertible.
So \<open>(r\<epsilon>)(\<rho>f\<^sup>*) = \<nu>\<^sup>-\<^sup>1(g\<gamma>)\<close> is invertible.''
\end{quotation}
We carry out the indicated calculations, inserting where required the canonical
isomorphisms omitted by CKS. It is perhaps amusing to compare the four-line sketch
given by CKS with the formalization below, but note that we have carried out the
proof in full, with no hand waving about units or associativities.
\<close>
have "section (g \<star> \<gamma>)"
proof
have "(g \<star> \<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w] \<cdot> (\<eta> \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[w]) \<cdot> (g \<star> \<gamma>) = g \<star> f\<^sup>*"
proof -
have "(\<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w] \<cdot> (\<eta> \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[w]) \<cdot> \<gamma> = f\<^sup>*"
proof -
have "(\<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w] \<cdot> (\<eta> \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[w]) \<cdot> \<gamma> =
(\<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w] \<cdot> (\<eta> \<star> w)) \<cdot> \<l>\<^sup>-\<^sup>1[w] \<cdot> \<gamma>"
using comp_assoc by auto
also have "... = (\<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w]) \<cdot> ((\<eta> \<star> w) \<cdot> (trg w \<star> \<gamma>)) \<cdot> \<l>\<^sup>-\<^sup>1[f\<^sup>*]"
using \<gamma> trg_\<gamma>_eq lunit'_naturality [of \<gamma>] comp_assoc by auto
also have "... = \<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> (\<a>[f\<^sup>*, f, w] \<cdot> ((f\<^sup>* \<star> f) \<star> \<gamma>)) \<cdot> (\<eta> \<star> f\<^sup>*) \<cdot> \<l>\<^sup>-\<^sup>1[f\<^sup>*]"
proof -
have "(\<eta> \<star> w) \<cdot> (trg w \<star> \<gamma>) = \<eta> \<star> \<gamma>"
using A \<gamma> interchange comp_arr_dom comp_cod_arr
by (metis T0.unit_simps(1-2) comp_ide_arr seqI' uw\<theta> w_in_hom(2) w_simps(4))
also have "... = ((f\<^sup>* \<star> f) \<star> \<gamma>) \<cdot> (\<eta> \<star> f\<^sup>*)"
using \<gamma> interchange comp_arr_dom comp_cod_arr T0.antipar T0.unit_simps(1,3)
in_homE
by metis
finally show ?thesis
using comp_assoc by simp
qed
also have "... = \<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> ((f\<^sup>* \<star> f \<star> \<gamma>) \<cdot> \<a>[f\<^sup>*, f, f\<^sup>*]) \<cdot> (\<eta> \<star> f\<^sup>*) \<cdot> \<l>\<^sup>-\<^sup>1[f\<^sup>*]"
using \<gamma> assoc_naturality [of "f\<^sup>*" f \<gamma>] trg_\<gamma>_eq T0.antipar by auto
also have "... = \<r>[f\<^sup>*] \<cdot> ((f\<^sup>* \<star> \<epsilon>) \<cdot> \<a>[f\<^sup>*, f, f\<^sup>*] \<cdot> (\<eta> \<star> f\<^sup>*)) \<cdot> \<l>\<^sup>-\<^sup>1[f\<^sup>*]"
using \<gamma> whisker_left trg_\<gamma>_eq T0.antipar comp_assoc by auto
also have "... = \<r>[f\<^sup>*] \<cdot> (\<r>\<^sup>-\<^sup>1[f\<^sup>*] \<cdot> \<l>[f\<^sup>*]) \<cdot> \<l>\<^sup>-\<^sup>1[f\<^sup>*]"
using T0.triangle_right by simp
also have "... = f\<^sup>*"
using comp_assoc by (simp add: comp_arr_dom comp_arr_inv')
finally show ?thesis by blast
qed
thus ?thesis
using \<gamma> whisker_left [of g "\<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w] \<cdot> (\<eta> \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[w]" \<gamma>]
T0.antipar
by simp
qed
thus "ide ((g \<star> \<r>[f\<^sup>*] \<cdot> (f\<^sup>* \<star> \<theta>') \<cdot> \<a>[f\<^sup>*, f, w] \<cdot> (\<eta> \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[w]) \<cdot> (g \<star> \<gamma>))"
using T0.antipar by simp
qed
moreover have "retraction (g \<star> \<gamma>)"
proof
have "\<guillemotleft>(g \<star> \<gamma>) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]) :
g \<star> w \<Rightarrow> g \<star> w\<guillemotright>"
using \<gamma> T0.antipar hseq_char by force
hence **: "arr ((g \<star> \<gamma>) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]))"
by auto
show "ide ((g \<star> \<gamma>) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]))"
proof -
have "((g \<star> \<gamma>) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])) =
g \<star> w"
proof -
have "((g \<star> \<gamma>) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])) =
\<nu> \<cdot> \<r>[r] \<cdot> ((r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<a>[src f\<^sup>*, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
(\<rho> \<star> trg w \<star> w)) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
proof -
have "(g \<star> \<gamma>) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]) =
(\<nu> \<cdot> \<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot> (\<rho> \<star> f\<^sup>*)) \<cdot>
(g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
using \<gamma> by auto
also have "... =
\<nu> \<cdot> \<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot>
((\<rho> \<star> f\<^sup>*) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>')) \<cdot>
(g \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
using comp_assoc by simp
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> (r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot>
(((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> ((r \<star> f) \<star> f\<^sup>* \<star> \<theta>') \<cdot> (\<rho> \<star> f\<^sup>* \<star> f \<star> w)) \<cdot>
(g \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (g \<star> \<eta> \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
proof -
have "(\<rho> \<star> f\<^sup>*) \<cdot> (g \<star> \<r>[f\<^sup>*]) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>') =
((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> (\<rho> \<star> f\<^sup>* \<star> src f\<^sup>*) \<cdot> (g \<star> f\<^sup>* \<star> \<theta>')"
proof -
have "(\<rho> \<star> f\<^sup>*) \<cdot> (g \<star> \<r>[f\<^sup>*]) = ((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> (\<rho> \<star> f\<^sup>* \<star> src f\<^sup>*)"
using tab_in_hom comp_arr_dom comp_cod_arr T0.antipar(1) interchange
by (metis T0.ide_right in_homE runit_simps(1,4-5))
thus ?thesis
by (metis comp_assoc)
qed
also have "... = ((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> (\<rho> \<star> f\<^sup>* \<star> \<theta>')"
using comp_arr_dom comp_cod_arr T0.antipar
interchange [of \<rho> g "f\<^sup>* \<star> src f\<^sup>*" "f\<^sup>* \<star> \<theta>'"]
by simp
also have "... = ((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> ((r \<star> f) \<star> f\<^sup>* \<star> \<theta>') \<cdot> (\<rho> \<star> f\<^sup>* \<star> f \<star> w)"
using comp_arr_dom comp_cod_arr T0.antipar
interchange [of "r \<star> f" \<rho> "f\<^sup>* \<star> \<theta>'" "f\<^sup>* \<star> f \<star> w"]
by simp
finally show ?thesis by simp
qed
also have "... =
\<nu> \<cdot> \<r>[r] \<cdot>
((r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot> ((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> ((r \<star> f) \<star> f\<^sup>* \<star> \<theta>')) \<cdot>
((\<rho> \<star> f\<^sup>* \<star> f \<star> w) \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (g \<star> \<eta> \<star> w)) \<cdot>
(g \<star> \<l>\<^sup>-\<^sup>1[w])"
using comp_assoc by simp
also have "... = \<nu> \<cdot> \<r>[r] \<cdot>
((r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<epsilon> \<star> f \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]) \<cdot>
(((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> ((r \<star> f) \<star> \<eta> \<star> w) \<cdot> (\<rho> \<star> trg w \<star> w)) \<cdot>
(g \<star> \<l>\<^sup>-\<^sup>1[w])"
proof -
have 1: "(r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot> ((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> ((r \<star> f) \<star> f\<^sup>* \<star> \<theta>') =
(r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<epsilon> \<star> f \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]"
proof -
have "(r \<star> \<epsilon>) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot> ((r \<star> f) \<star> \<r>[f\<^sup>*]) \<cdot> ((r \<star> f) \<star> f\<^sup>* \<star> \<theta>') =
(r \<star> \<epsilon>) \<cdot> (r \<star> f \<star> \<r>[f\<^sup>*]) \<cdot> \<a>[r, f, f\<^sup>* \<star> src f\<^sup>*] \<cdot> ((r \<star> f) \<star> f\<^sup>* \<star> \<theta>')"
proof -
have "\<a>[r, f, f\<^sup>*] \<cdot> ((r \<star> f) \<star> \<r>[f\<^sup>*]) = (r \<star> f \<star> \<r>[f\<^sup>*]) \<cdot> \<a>[r, f, f\<^sup>* \<star> src f\<^sup>*]"
using assoc_naturality [of r f "\<r>[f\<^sup>*]"] T0.antipar by auto
thus ?thesis
using comp_assoc by metis
qed
also have "... = (r \<star> \<epsilon>) \<cdot> (r \<star> f \<star> \<r>[f\<^sup>*]) \<cdot> (r \<star> f \<star> f\<^sup>* \<star> \<theta>') \<cdot>
\<a>[r, f, f\<^sup>* \<star> f \<star> w]"
using assoc_naturality [of r f "f\<^sup>* \<star> \<theta>'"] T0.antipar by fastforce
also have "... = (r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*]) \<cdot>
(r \<star> f \<star> f\<^sup>* \<star> \<theta>') \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]"
proof -
have "(r \<star> \<epsilon>) \<cdot> (r \<star> f \<star> \<r>[f\<^sup>*]) =
(r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*])"
proof -
have "(r \<star> \<epsilon>) \<cdot> (r \<star> f \<star> \<r>[f\<^sup>*]) = r \<star> (\<epsilon> \<cdot> (f \<star> \<r>[f\<^sup>*]))"
using whisker_left T0.antipar by simp
also have "... =
(r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*])"
proof -
have "\<epsilon> \<cdot> (f \<star> \<r>[f\<^sup>*]) = \<r>[src f\<^sup>*] \<cdot> (\<epsilon> \<star> src f\<^sup>*) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*]"
using ide_leg0 T0.antipar runit_hcomp invert_side_of_triangle(2)
runit_naturality comp_assoc
by (metis (no_types, lifting) T0.counit_simps(1-4) T0.ide_right)
thus ?thesis
using whisker_left T0.antipar by simp
qed
finally show ?thesis by simp
qed
thus ?thesis using comp_assoc by metis
qed
also have "... =
(r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*]) \<cdot> (r \<star> f \<star> f\<^sup>* \<star> \<theta>')) \<cdot>
\<a>[r, f, f\<^sup>* \<star> f \<star> w]"
using comp_assoc by simp
also have "... = (r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot>
((r \<star> (f \<star> f\<^sup>*) \<star> \<theta>') \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w])) \<cdot>
\<a>[r, f, f\<^sup>* \<star> f \<star> w]"
proof -
have "(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*]) \<cdot> (r \<star> f \<star> f\<^sup>* \<star> \<theta>') =
(r \<star> (f \<star> f\<^sup>*) \<star> \<theta>') \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w])"
proof -
have "(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*]) \<cdot> (r \<star> f \<star> f\<^sup>* \<star> \<theta>') =
r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, src f\<^sup>*] \<cdot> (f \<star> f\<^sup>* \<star> \<theta>')"
using whisker_left T0.antipar by simp
also have "... = r \<star> ((f \<star> f\<^sup>*) \<star> \<theta>') \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]"
using assoc'_naturality [of f "f\<^sup>*" \<theta>'] T0.antipar by auto
also have "... = (r \<star> (f \<star> f\<^sup>*) \<star> \<theta>') \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w])"
using whisker_left T0.antipar by auto
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = (r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> (f \<star> f\<^sup>*) \<star> \<theta>') \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]"
using comp_assoc by simp
also have "... =
(r \<star> \<r>[src f\<^sup>*]) \<cdot> ((r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> (f \<star> f\<^sup>*) \<star> \<theta>')) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]"
using comp_assoc by simp
also have "... = (r \<star> \<r>[src f\<^sup>*]) \<cdot> ((r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<epsilon> \<star> f \<star> w)) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]"
proof -
have "(r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> (f \<star> f\<^sup>*) \<star> \<theta>') =
(r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<epsilon> \<star> f \<star> w)"
proof -
have "(r \<star> \<epsilon> \<star> src f\<^sup>*) \<cdot> (r \<star> (f \<star> f\<^sup>*) \<star> \<theta>') =
r \<star> (\<epsilon> \<star> src f\<^sup>*) \<cdot> ((f \<star> f\<^sup>*) \<star> \<theta>')"
using whisker_left T0.antipar by simp
also have "... = r \<star> \<epsilon> \<star> \<theta>'"
using interchange [of \<epsilon> "f \<star> f\<^sup>*" "src f\<^sup>*" \<theta>']
T0.antipar comp_arr_dom comp_cod_arr
by auto
also have "... = r \<star> (src f\<^sup>* \<star> \<theta>') \<cdot> (\<epsilon> \<star> f \<star> w)"
using interchange [of "src f\<^sup>*" \<epsilon> \<theta>' "f \<star> w"]
T0.antipar comp_arr_dom comp_cod_arr
by auto
also have "... = (r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<epsilon> \<star> f \<star> w)"
using whisker_left T0.antipar by simp
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = (r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<epsilon> \<star> f \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]"
using comp_assoc by simp
finally show ?thesis by simp
qed
have 2: "(\<rho> \<star> f\<^sup>* \<star> f \<star> w) \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (g \<star> \<eta> \<star> w) =
((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> ((r \<star> f) \<star> \<eta> \<star> w) \<cdot> (\<rho> \<star> trg w \<star> w)"
proof -
have "(\<rho> \<star> f\<^sup>* \<star> f \<star> w) \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (g \<star> \<eta> \<star> w) =
((\<rho> \<star> f\<^sup>* \<star> f \<star> w) \<cdot> (g \<star> \<a>[f\<^sup>*, f, w])) \<cdot> (g \<star> \<eta> \<star> w)"
using comp_assoc by simp
also have "... = (((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (\<rho> \<star> (f\<^sup>* \<star> f) \<star> w)) \<cdot> (g \<star> \<eta> \<star> w)"
proof -
have "(\<rho> \<star> f\<^sup>* \<star> f \<star> w) \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) =
((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (\<rho> \<star> (f\<^sup>* \<star> f) \<star> w)"
proof -
have "(\<rho> \<star> f\<^sup>* \<star> f \<star> w) \<cdot> (g \<star> \<a>[f\<^sup>*, f, w]) =
\<rho> \<cdot> g \<star> (f\<^sup>* \<star> f \<star> w) \<cdot> \<a>[f\<^sup>*, f, w]"
using interchange T0.antipar by auto
also have "... = \<rho> \<star> \<a>[f\<^sup>*, f, w]"
using comp_arr_dom comp_cod_arr T0.antipar by auto
also have "... = (r \<star> f) \<cdot> \<rho> \<star> \<a>[f\<^sup>*, f, w] \<cdot> ((f\<^sup>* \<star> f) \<star> w)"
using comp_arr_dom comp_cod_arr T0.antipar by auto
also have "... = ((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (\<rho> \<star> (f\<^sup>* \<star> f) \<star> w)"
using interchange T0.antipar by auto
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = ((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> (\<rho> \<star> (f\<^sup>* \<star> f) \<star> w) \<cdot> (g \<star> \<eta> \<star> w)"
using comp_assoc by simp
also have "... = ((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot> ((r \<star> f) \<star> \<eta> \<star> w) \<cdot> (\<rho> \<star> trg w \<star> w)"
proof -
have "(\<rho> \<star> (f\<^sup>* \<star> f) \<star> w) \<cdot> (g \<star> \<eta> \<star> w) = ((r \<star> f) \<star> \<eta> \<star> w) \<cdot> (\<rho> \<star> trg w \<star> w)"
proof -
have "(\<rho> \<star> (f\<^sup>* \<star> f) \<star> w) \<cdot> (g \<star> \<eta> \<star> w) = \<rho> \<cdot> g \<star> (f\<^sup>* \<star> f) \<cdot> \<eta> \<star> w \<cdot> w"
proof -
have "\<guillemotleft>g \<star> \<eta> \<star> w : g \<star> trg w \<star> w \<Rightarrow> g \<star> (f\<^sup>* \<star> f) \<star> w\<guillemotright>"
by (intro hcomp_in_vhom, auto)
thus ?thesis
using interchange whisker_right T0.antipar by auto
qed
also have "... = (r \<star> f) \<cdot> \<rho> \<star> \<eta> \<cdot> trg w \<star> w \<cdot> w"
using comp_arr_dom comp_cod_arr by auto
also have "... = ((r \<star> f) \<star> \<eta> \<star> w) \<cdot> (\<rho> \<star> trg w \<star> w)"
using interchange [of "r \<star> f" \<rho> "\<eta> \<star> w" "trg w \<star> w"]
interchange [of \<eta> "trg w" w w]
comp_arr_dom comp_cod_arr T0.unit_in_hom
by auto
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
show ?thesis
using 1 2 by simp
qed
also have "... =
\<nu> \<cdot> \<r>[r] \<cdot>
((r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>') \<cdot>
((r \<star> \<a>[src r, f, w]) \<cdot> (r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f \<star> f\<^sup>*, f, w])) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot> \<a>[r, f, f\<^sup>* \<star> f \<star> w]) \<cdot>
(((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
(\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot> (r \<star> \<a>[f, f\<^sup>* \<star> f, w]) \<cdot>
(r \<star> (f \<star> \<eta>) \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w]) \<cdot>
(\<rho> \<star> trg w \<star> w)) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
proof -
have 3: "r \<star> \<epsilon> \<star> f \<star> w =
(r \<star> \<a>[src r, f, w]) \<cdot> (r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f \<star> f\<^sup>*, f, w])"
proof -
have "r \<star> \<epsilon> \<star> f \<star> w =
((r \<star> \<a>[src r, f, w]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[src r, f, w])) \<cdot> (r \<star> \<epsilon> \<star> f \<star> w)"
using T0.antipar whisker_left [of r "\<a>[src r, f, w]" "\<a>\<^sup>-\<^sup>1[src r, f, w]"]
comp_cod_arr comp_assoc_assoc'
by simp
also have "... = (r \<star> \<a>[src r, f, w]) \<cdot> (r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f \<star> f\<^sup>*, f, w])"
using assoc'_naturality [of \<epsilon> f w]
whisker_left [of r "\<a>\<^sup>-\<^sup>1[src r, f, w]" "\<epsilon> \<star> f \<star> w"]
whisker_left comp_assoc T0.antipar
by simp
finally show ?thesis
using T0.antipar by simp
qed
have 4: "(r \<star> f) \<star> \<eta> \<star> w =
\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot> (r \<star> \<a>[f, f\<^sup>* \<star> f, w]) \<cdot>
(r \<star> (f \<star> \<eta>) \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w]"
proof -
have "(r \<star> f) \<star> \<eta> \<star> w =
(\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot>
((r \<star> \<a>[f, f\<^sup>* \<star> f, w]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>* \<star> f, w])) \<cdot>
\<a>[r, f, (f\<^sup>* \<star> f) \<star> w]) \<cdot>
((r \<star> f) \<star> \<eta> \<star> w)"
proof -
have "ide r" by simp
moreover have "seq \<a>[f, f\<^sup>* \<star> f, w] \<a>\<^sup>-\<^sup>1[f, f\<^sup>* \<star> f, w]"
using T0.antipar comp_cod_arr ide_base by simp
ultimately have "(r \<star> \<a>[f, f\<^sup>* \<star> f, w]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>* \<star> f, w]) =
r \<star> \<a>[f, f\<^sup>* \<star> f, w] \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sup>* \<star> f, w]"
using whisker_left by metis
thus ?thesis
using T0.antipar comp_cod_arr comp_assoc_assoc' by simp
qed
also have "... =
\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot>
(r \<star> \<a>[f, f\<^sup>* \<star> f, w]) \<cdot> ((r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>* \<star> f, w]) \<cdot>
(r \<star> f \<star> \<eta> \<star> w)) \<cdot>
\<a>[r, f, trg w \<star> w]"
using assoc_naturality [of r f "\<eta> \<star> w"] comp_assoc by fastforce
also have "... =
\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot>
(r \<star> \<a>[f, f\<^sup>* \<star> f, w]) \<cdot> (r \<star> (f \<star> \<eta>) \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
\<a>[r, f, trg w \<star> w]"
using assoc'_naturality [of f \<eta> w] T0.antipar comp_assoc
whisker_left [of r "\<a>\<^sup>-\<^sup>1[f, f\<^sup>* \<star> f, w]" "f \<star> \<eta> \<star> w"]
whisker_left [of r "(f \<star> \<eta>) \<star> w" "\<a>\<^sup>-\<^sup>1[f, trg w, w]"]
by simp
finally show ?thesis by blast
qed
show ?thesis
using 3 4 T0.antipar by simp
qed
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> ((r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>') \<cdot>
(r \<star> \<a>[src r, f, w]) \<cdot>
((r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[f \<star> f\<^sup>*, f, w]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot>
\<a>[r, f, f\<^sup>* \<star> f \<star> w] \<cdot> ((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot> (r \<star> \<a>[f, f\<^sup>* \<star> f, w])) \<cdot>
(r \<star> (f \<star> \<eta>) \<star> w)) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
(\<rho> \<star> trg w \<star> w)) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
using comp_assoc T0.antipar by auto
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> ((r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>') \<cdot>
(r \<star> \<a>[src r, f, w]) \<cdot>
((r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<star> w) \<cdot>
(r \<star> (f \<star> \<eta>) \<star> w)) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
(\<rho> \<star> trg w \<star> w)) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
proof -
have "(r \<star> \<a>\<^sup>-\<^sup>1[f \<star> f\<^sup>*, f, w]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot>
\<a>[r, f, f\<^sup>* \<star> f \<star> w] \<cdot> ((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot> (r \<star> \<a>[f, f\<^sup>* \<star> f, w]) =
r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<star> w"
proof -
text \<open>We can compress the reasoning about the associativities using coherence.\<close>
have "(r \<star> \<a>\<^sup>-\<^sup>1[f \<star> f\<^sup>*, f, w]) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f \<star> w]) \<cdot>
\<a>[r, f, f\<^sup>* \<star> f \<star> w] \<cdot> ((r \<star> f) \<star> \<a>[f\<^sup>*, f, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[r, f, (f\<^sup>* \<star> f) \<star> w] \<cdot> (r \<star> \<a>[f, f\<^sup>* \<star> f, w]) =
\<lbrace>(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> ((\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>f\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>])\<rbrace>"
using T0.antipar \<a>'_def \<alpha>_def assoc'_eq_inv_assoc by auto
also have "... = \<lbrace>\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<rbrace>"
using T0.antipar by (intro E.eval_eqI, auto)
also have "... = r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<star> w"
using T0.antipar \<a>'_def \<alpha>_def assoc'_eq_inv_assoc by simp
finally show ?thesis
by simp
qed
thus ?thesis by simp
qed
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> ((r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>') \<cdot>
(r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
(\<rho> \<star> trg w \<star> w)) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
proof -
have "(r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<star> w) \<cdot> (r \<star> (f \<star> \<eta>) \<star> w) =
r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w"
proof -
have "(r \<star> (\<epsilon> \<star> f) \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<star> w) \<cdot> (r \<star> (f \<star> \<eta>) \<star> w) =
r \<star> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, f\<^sup>*, f] \<cdot> (f \<star> \<eta>) \<star> w"
using whisker_left whisker_right T0.antipar by simp
also have "... = r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w"
using T0.triangle_left by simp
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> ((r \<star> \<r>[src f\<^sup>*]) \<cdot> (r \<star> src f\<^sup>* \<star> \<theta>') \<cdot> (r \<star> \<a>[src f\<^sup>*, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
(\<rho> \<star> trg w \<star> w)) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w])"
using T0.antipar by simp
finally show ?thesis by simp
qed
also have "... = \<nu> \<cdot> \<r>[r] \<cdot>
((r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>')) \<cdot>
(r \<star> \<a>[src r, f, w]) \<cdot> (r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
((\<rho> \<star> trg w \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]))"
using comp_assoc T0.antipar by simp
also have "... = \<nu> \<cdot> \<r>[r] \<cdot>
((r \<star> \<theta>') \<cdot> (r \<star> \<l>[f \<star> w])) \<cdot>
(r \<star> \<a>[src r, f, w]) \<cdot> (r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot>
(r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot> \<a>[r, f, trg w \<star> w] \<cdot>
(((r \<star> f) \<star> \<l>\<^sup>-\<^sup>1[w]) \<cdot> (\<rho> \<star> w))"
proof -
have "(r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>') = (r \<star> \<theta>') \<cdot> (r \<star> \<l>[f \<star> w])"
proof -
have "(r \<star> \<r>[src r]) \<cdot> (r \<star> src r \<star> \<theta>') = r \<star> \<r>[src r] \<cdot> (src r \<star> \<theta>')"
using whisker_left by simp
also have "... = r \<star> \<theta>' \<cdot> \<l>[f \<star> w]"
using lunit_naturality [of \<theta>'] unitor_coincidence by simp
also have "... = (r \<star> \<theta>') \<cdot> (r \<star> \<l>[f \<star> w])"
using whisker_left by simp
finally show ?thesis by simp
qed
moreover have "(\<rho> \<star> trg w \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]) = ((r \<star> f) \<star> \<l>\<^sup>-\<^sup>1[w]) \<cdot> (\<rho> \<star> w)"
proof -
have "(\<rho> \<star> trg w \<star> w) \<cdot> (g \<star> \<l>\<^sup>-\<^sup>1[w]) = \<rho> \<cdot> g \<star> (trg w \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[w]"
using interchange by simp
also have "... = \<rho> \<star> \<l>\<^sup>-\<^sup>1[w]"
using comp_arr_dom comp_cod_arr by simp
also have "... = (r \<star> f) \<cdot> \<rho> \<star> \<l>\<^sup>-\<^sup>1[w] \<cdot> w"
using comp_arr_dom comp_cod_arr by simp
also have "... = ((r \<star> f) \<star> \<l>\<^sup>-\<^sup>1[w]) \<cdot> (\<rho> \<star> w)"
using interchange by simp
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> (r \<star> \<theta>') \<cdot>
((r \<star> \<l>[f \<star> w]) \<cdot> (r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
\<a>[r, f, trg w \<star> w] \<cdot> ((r \<star> f) \<star> \<l>\<^sup>-\<^sup>1[w])) \<cdot>
(\<rho> \<star> w)"
using comp_assoc by simp
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> (r \<star> \<theta>') \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w)"
proof -
have "((r \<star> \<l>[f \<star> w]) \<cdot> (r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
\<a>[r, f, trg w \<star> w] \<cdot> ((r \<star> f) \<star> \<l>\<^sup>-\<^sup>1[w])) =
\<a>[r, f, w]"
proof -
have "((r \<star> \<l>[f \<star> w]) \<cdot> (r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
\<a>[r, f, trg w \<star> w] \<cdot> ((r \<star> f) \<star> \<l>\<^sup>-\<^sup>1[w])) =
((r \<star> (\<l>[f] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, f, w]) \<cdot> (r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
(r \<star> f \<star> \<l>\<^sup>-\<^sup>1[w])) \<cdot> \<a>[r, f, w]"
using comp_assoc assoc_naturality [of r f "\<l>\<^sup>-\<^sup>1[w]"] lunit_hcomp by simp
also have "... = \<a>[r, f, w]"
proof -
have "(r \<star> (\<l>[f] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, f, w]) \<cdot> (r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
(r \<star> f \<star> \<l>\<^sup>-\<^sup>1[w]) =
r \<star> f \<star> w"
proof -
text \<open>Again, get a little more mileage out of coherence.\<close>
have "(r \<star> (\<l>[f] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[trg f, f, w]) \<cdot> (r \<star> \<a>[src r, f, w]) \<cdot>
(r \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, trg w, w]) \<cdot>
(r \<star> f \<star> \<l>\<^sup>-\<^sup>1[w]) =
\<lbrace>(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<l>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>) \<^bold>\<cdot> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[E.Trg \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[E.Src \<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<r>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>) \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, E.Trg \<^bold>\<langle>w\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>w\<^bold>\<rangle>\<^bold>])\<rbrace>"
using \<ll>_ide_simp \<rr>_ide_simp \<a>'_def \<alpha>_def assoc'_eq_inv_assoc by simp
also have "... = \<lbrace>\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<rbrace>"
by (intro E.eval_eqI, auto)
also have "... = r \<star> f \<star> w"
by simp
finally show ?thesis by blast
qed
thus ?thesis
using comp_cod_arr
by (metis assoc_is_natural_1 base_simps(2-3) leg0_simps(2-4)
w_simps(2) w_simps(4) w_simps(5))
qed
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = \<nu> \<cdot> \<r>[r] \<cdot> \<r>\<^sup>-\<^sup>1[r] \<cdot> inv \<nu>"
proof -
have "\<r>\<^sup>-\<^sup>1[r] \<cdot> inv \<nu> = (r \<star> \<theta>') \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w)"
using ** w\<theta>'\<nu> ide_base ide_leg0 tab_in_hom invert_side_of_triangle(2) comp_arr_dom
T0.antipar comp_assoc runit'_simps(1)
by metis
thus ?thesis by simp
qed
also have "... = g \<star> w"
using ** w\<theta>'\<nu> ide_base comp_arr_inv'
by (metis calculation in_homE invert_side_of_triangle(1) iso_runit iso_runit')
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
qed
ultimately have 1: "iso (g \<star> \<gamma>)"
using iso_iff_section_and_retraction by simp
have "iso (inv (\<nu> \<cdot> \<r>[r]) \<cdot> (g \<star> \<gamma>))"
proof -
have "iso (inv (\<nu> \<cdot> \<r>[r]))"
using w\<theta>'\<nu> \<gamma> iso_runit
by (elim conjE in_homE, intro iso_inv_iso isos_compose, auto)
thus ?thesis
using 1 w\<theta>'\<nu> \<gamma> trg_\<gamma>_eq isos_compose
by (elim conjE in_homE, auto)
qed
moreover have "inv (\<nu> \<cdot> \<r>[r]) \<cdot> (g \<star> \<gamma>) = composite_cell f\<^sup>* \<epsilon>"
proof -
have "inv (\<nu> \<cdot> \<r>[r]) \<cdot> (g \<star> \<gamma>) = inv (\<nu> \<cdot> \<r>[r]) \<cdot> \<nu> \<cdot> \<r>[r] \<cdot> composite_cell f\<^sup>* \<epsilon>"
using \<gamma> by auto
also have "... = ((inv (\<nu> \<cdot> \<r>[r]) \<cdot> (\<nu> \<cdot> \<r>[r])) \<cdot> (r \<star> \<epsilon>)) \<cdot> \<a>[r, f, f\<^sup>*] \<cdot> (\<rho> \<star> f\<^sup>*)"
using w\<theta>'\<nu> comp_assoc by auto
also have "... = composite_cell f\<^sup>* \<epsilon>"
proof -
have "dom \<nu> = r"
using w\<theta>'\<nu> by auto
thus ?thesis
using iso_runit w\<theta>'\<nu> isos_compose comp_cod_arr whisker_left comp_inv_arr'
by auto
qed
finally show ?thesis by blast
qed
ultimately have "iso (composite_cell f\<^sup>* \<epsilon>)" by simp
thus "iso (T0.trnr\<^sub>\<epsilon> r \<rho>)"
using T0.trnr\<^sub>\<epsilon>_def ide_base runit_in_hom iso_runit isos_compose
by (metis A arrI seqE)
qed
text \<open>
It is convenient to have a simpler version of the previous result for when we do
not care about the details of the isomorphism.
\<close>
lemma yields_isomorphic_representation':
obtains \<psi> where "\<guillemotleft>\<psi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>" and "iso \<psi>"
using yields_isomorphic_representation adjoint_pair_def by simp
end
text \<open>
It is natural to ask whether if \<open>\<guillemotleft>\<psi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>\<close> is an isomorphism
then \<open>\<rho> = (\<psi> \<star> f) \<cdot> T0.trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)\<close> is a tabulation of \<open>r\<close>.
This is not true without additional conditions on \<open>f\<close> and \<open>g\<close>
(\emph{cf.}~the comments following CKS Proposition 6).
So only rather special isomorphisms \<open>\<guillemotleft>\<psi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright>\<close> result from tabulations of \<open>r\<close>.
\<close>
subsection "Tabulation of a Right Adjoint"
text \<open>
Here we obtain a tabulation of the right adjoint of a map. This is CKS Proposition 1(e).
It was somewhat difficult to find the correct way to insert the unitors
that CKS omit. At first I thought I could only prove this under the assumption
that the bicategory is normal, but later I saw how to do it in the general case.
\<close>
context adjunction_in_bicategory
begin
lemma tabulation_of_right_adjoint:
shows "tabulation V H \<a> \<i> src trg g \<eta> f (src f)"
proof -
interpret T: tabulation_data V H \<a> \<i> src trg g \<eta> f \<open>src f\<close>
using unit_in_hom antipar by (unfold_locales, simp_all)
show ?thesis
proof
show T1: "\<And>u \<omega>. \<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> g \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> src f \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
proof -
fix u v \<omega>
assume u: "ide u"
assume \<omega>: "\<guillemotleft>\<omega> : v \<Rightarrow> g \<star> u\<guillemotright>"
have v: "ide v"
using \<omega> by auto
have 1: "src g = trg u"
using \<omega> by (metis arr_cod in_homE not_arr_null seq_if_composable)
have 2: "src f = trg v"
using \<omega> 1 u ide_right antipar(1) vconn_implies_hpar(4) by force
text \<open>It seems clear that we need to take \<open>w = v\<close> and \<open>\<nu> = \<l>\<^sup>-\<^sup>1[v]\<close>. \<close>
let ?w = v
let ?\<nu> = "\<l>\<^sup>-\<^sup>1[v]"
have \<nu>: "\<guillemotleft>?\<nu> : v \<Rightarrow> src f \<star> ?w\<guillemotright> \<and> iso ?\<nu>"
using v 2 iso_lunit' by auto
text \<open>
We need \<open>\<theta>\<close>, defined to satisfy \<open>\<guillemotleft>\<theta> : f \<star> v \<Rightarrow> u\<guillemotright>\<close> and
\<open>\<omega> = (v \<star> \<theta>) \<cdot> \<a>[v, f, v] \<cdot> (\<eta> \<star> w) \<cdot> \<l>\<^sup>-\<^sup>1[v]\<close>.
We have \<open>\<guillemotleft>\<omega> : v \<Rightarrow> g \<star> u\<guillemotright>\<close>, so we can get arrow \<open>\<guillemotleft>\<theta> : f \<star> v \<Rightarrow> u\<guillemotright>\<close> by adjoint transpose.
Note that this uses adjoint transpose on the \emph{left}, rather than on the right.
\<close>
let ?\<theta> = "trnl\<^sub>\<epsilon> u \<omega>"
have \<theta>: "\<guillemotleft>?\<theta> : f \<star> ?w \<Rightarrow> u\<guillemotright>"
using u v antipar 1 2 \<omega> adjoint_transpose_left(2) [of u v] by auto
text \<open>
Now, \<open>trnl\<^sub>\<eta> v \<theta> \<equiv> (g \<star> \<theta>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]\<close>, which suggests that
we ought to have \<open>\<omega> = trnl\<^sub>\<eta> v \<theta>\<close> and \<open>\<nu> = \<l>\<^sup>-\<^sup>1[v]\<close>;
\<close>
have "T.composite_cell ?w ?\<theta> \<cdot> ?\<nu> = \<omega>"
using u v \<omega> 1 2 adjoint_transpose_left(4) [of u v \<omega>] trnl\<^sub>\<eta>_def comp_assoc by simp
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> src f \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
using v \<theta> \<nu> antipar comp_assoc by blast
qed
show T2: "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w'; \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : src f \<star> w \<Rightarrow> src f \<star> w'\<guillemotright>;
T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume w: "ide w"
assume w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : src f \<star> w \<Rightarrow> src f \<star> w'\<guillemotright>"
assume E: "T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<cdot> \<beta>"
interpret T: uw\<theta>w'\<theta>'\<beta> V H \<a> \<i> src trg g \<eta> f \<open>src f\<close> u w \<theta> w' \<theta>' \<beta>
using w w' \<theta> \<theta>' \<beta> E comp_assoc by (unfold_locales, auto)
have 2: "src f = trg \<beta>"
using antipar by simp
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
proof -
text \<open>
The requirement \<open>\<beta> = src f \<star> \<gamma>\<close> means we have to essentially invert \<open>\<lambda>\<gamma>. src f \<star> \<gamma>\<close>
to obtain \<open>\<gamma>\<close>. CKS say only: ``the strong form of \<open>T2\<close> is clear since \<open>g = 1\<close>"
(here by ``\<open>g\<close>'' they are referring to \<open>dom \<eta>\<close>, the ``output leg'' of the span in
the tabulation). This would mean that we would have to take \<open>\<gamma> = \<beta>\<close>, which doesn't
work for a general bicategory (we don't necessarily have \<open>src f \<star> \<gamma> = \<gamma>\<close>).
For a general bicategory, we have to take \<open>\<gamma> = \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w]\<close>.
\<close>
let ?\<gamma> = "\<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w]"
have \<gamma>: "\<guillemotleft>?\<gamma> : w \<Rightarrow> w'\<guillemotright>"
using \<beta> by simp
have 3: "\<beta> = src f \<star> ?\<gamma>"
proof -
have "\<beta> = \<l>\<^sup>-\<^sup>1[w'] \<cdot> ?\<gamma> \<cdot> \<l>[w]"
using \<beta> iso_lunit
by (simp add: comp_arr_dom invert_side_of_triangle(1) comp_assoc)
also have "... = \<l>\<^sup>-\<^sup>1[w'] \<cdot> \<l>[w'] \<cdot> (src f \<star> ?\<gamma>)"
using \<gamma> lunit_naturality
by (metis T.uw\<theta>.w_simps(4) in_homE trg_dom)
also have "... = (\<l>\<^sup>-\<^sup>1[w'] \<cdot> \<l>[w']) \<cdot> (src f \<star> ?\<gamma>)"
using comp_assoc by simp
also have "... = src f \<star> ?\<gamma>"
using \<gamma> iso_lunit comp_inv_arr comp_cod_arr
by (metis T.\<beta>_simps(1) calculation comp_ide_arr inv_is_inverse inverse_arrowsE w')
finally show ?thesis by simp
qed
have "\<theta> = \<theta>' \<cdot> (f \<star> ?\<gamma>)"
proof -
have "\<theta> = trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> w \<theta>)"
using \<theta> adjoint_transpose_left(3) [of u w \<theta>] by simp
also have "... = trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> w' \<theta>' \<cdot> \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w])"
proof -
have "trnl\<^sub>\<eta> w \<theta> = trnl\<^sub>\<eta> w' \<theta>' \<cdot> \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w]"
proof -
have "trnl\<^sub>\<eta> w \<theta> \<cdot> \<l>[w] = (T.composite_cell w \<theta> \<cdot> \<l>\<^sup>-\<^sup>1[w]) \<cdot> \<l>[w]"
unfolding trnl\<^sub>\<eta>_def using comp_assoc by simp
also have "... = T.composite_cell w \<theta> \<cdot> (\<l>\<^sup>-\<^sup>1[w] \<cdot> \<l>[w])"
using comp_assoc by simp
also have 4: "... = T.composite_cell w \<theta>"
using comp_arr_dom by (simp add: comp_inv_arr')
also have "... = T.composite_cell w' \<theta>' \<cdot> \<beta>"
using E by simp
also have "... = (T.composite_cell w' \<theta>' \<cdot> \<l>\<^sup>-\<^sup>1[w']) \<cdot> \<l>[w'] \<cdot> \<beta>"
proof -
have "(\<l>\<^sup>-\<^sup>1[w'] \<cdot> \<l>[w']) \<cdot> \<beta> = \<beta>"
using iso_lunit \<beta> comp_cod_arr comp_assoc comp_inv_arr' by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = trnl\<^sub>\<eta> w' \<theta>' \<cdot> \<l>[w'] \<cdot> \<beta>"
unfolding trnl\<^sub>\<eta>_def using comp_assoc by simp
finally have "trnl\<^sub>\<eta> w \<theta> \<cdot> \<l>[w] = trnl\<^sub>\<eta> w' \<theta>' \<cdot> \<l>[w'] \<cdot> \<beta>"
by simp
thus ?thesis
using \<beta> 4 invert_side_of_triangle(2) adjoint_transpose_left iso_lunit
trnl\<^sub>\<eta>_def comp_assoc
by metis
qed
thus ?thesis by simp
qed
also have "... = \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> trnl\<^sub>\<eta> w' \<theta>' \<cdot> \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w])"
using trnl\<^sub>\<epsilon>_def by simp
also have
"... = \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> trnl\<^sub>\<eta> w' \<theta>') \<cdot> (f \<star> \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w])"
using ide_left ide_right w w' 2 \<beta> \<theta> antipar trnl\<^sub>\<epsilon>_def adjoint_transpose_left
whisker_left
by (metis T.uw\<theta>.\<theta>_simps(1) calculation hseqE seqE)
also have
"... = (\<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> trnl\<^sub>\<eta> w' \<theta>')) \<cdot> (f \<star> \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w])"
using comp_assoc by simp
also have "... = trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> w' \<theta>') \<cdot> (f \<star> \<l>[w'] \<cdot> \<beta> \<cdot> \<l>\<^sup>-\<^sup>1[w])"
unfolding trnl\<^sub>\<epsilon>_def by simp
also have "... = \<theta>' \<cdot> (f \<star> ?\<gamma>)"
using \<theta>' adjoint_transpose_left(3) by auto
finally show ?thesis by simp
qed
hence "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
using \<gamma> 3 hcomp_obj_arr by blast
moreover have "\<And>\<gamma> \<gamma>'. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>) \<and>
\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>') \<Longrightarrow> \<gamma> = \<gamma>'"
proof -
fix \<gamma> \<gamma>'
assume \<gamma>\<gamma>': "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>) \<and>
\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = src f \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>')"
show "\<gamma> = \<gamma>'"
using \<gamma>\<gamma>' vconn_implies_hpar(2) L.is_faithful [of \<gamma> \<gamma>'] by force
qed
ultimately show ?thesis by blast
qed
qed
qed
qed
end
subsection "Preservation by Isomorphisms"
text \<open>
Next, we show that tabulations are preserved under composition on all three sides by
isomorphisms. This is something that we would expect to hold if ``tabulation'' is a
properly bicategorical notion.
\<close>
context tabulation
begin
text \<open>
Tabulations are preserved under composition of an isomorphism with the ``input leg''.
\<close>
lemma preserved_by_input_iso:
assumes "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "tabulation V H \<a> \<i> src trg r ((r \<star> \<phi>) \<cdot> \<rho>) f' g"
proof -
interpret T': tabulation_data V H \<a> \<i> src trg r \<open>(r \<star> \<phi>) \<cdot> \<rho>\<close> f'
using assms(1) tab_in_hom
apply unfold_locales
apply auto
by force
show ?thesis
proof
show "\<And>u \<omega>. \<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and>
iso \<nu> \<and> T'.composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
proof -
fix u \<omega>
assume u: "ide u" and \<omega>: "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
obtain w \<theta> \<nu> where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and>
iso \<nu> \<and> composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
using u \<omega> T1 by blast
interpret T1: uw\<theta>\<omega>\<nu> V H \<a> \<i> src trg r \<rho> f g u w \<theta> \<omega> \<nu>
using w\<theta>\<nu> comp_assoc by (unfold_locales, auto)
have 1: "\<guillemotleft>inv \<phi> \<star> w : f' \<star> w \<Rightarrow> f \<star> w\<guillemotright>"
using assms by (intro hcomp_in_vhom, auto)
have "ide w \<and> \<guillemotleft>\<theta> \<cdot> (inv \<phi> \<star> w) : f' \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T'.composite_cell w (\<theta> \<cdot> (inv \<phi> \<star> w)) \<cdot> \<nu> = \<omega>"
using w\<theta>\<nu> 1
apply (intro conjI)
apply auto[4]
proof -
show "T'.composite_cell w (\<theta> \<cdot> (inv \<phi> \<star> w)) \<cdot> \<nu> = \<omega>"
proof -
have "T'.composite_cell w (\<theta> \<cdot> (inv \<phi> \<star> w)) \<cdot> \<nu> =
(r \<star> \<theta>) \<cdot> ((r \<star> inv \<phi> \<star> w) \<cdot> \<a>[r, f', w]) \<cdot> ((r \<star> \<phi>) \<cdot> \<rho> \<star> w) \<cdot> \<nu>"
using assms(1) 1 whisker_left [of r \<theta> "inv \<phi> \<star> w"] comp_assoc by auto
also have "... = (r \<star> \<theta>) \<cdot> (\<a>[r, f, w] \<cdot> ((r \<star> inv \<phi>) \<star> w)) \<cdot> ((r \<star> \<phi>) \<cdot> \<rho> \<star> w) \<cdot> \<nu>"
using assms assoc_naturality [of r "inv \<phi>" w]
by (metis 1 T'.tab_simps(1) base_simps(3) base_simps(4) T1.w_simps(5-6)
cod_inv dom_inv hseqE in_homE seqE trg_inv)
also have "... = (r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> ((((r \<star> inv \<phi>) \<star> w) \<cdot> ((r \<star> \<phi>) \<star> w)) \<cdot> (\<rho> \<star> w)) \<cdot> \<nu>"
using whisker_right [of w "r \<star> \<phi>" \<rho>] comp_assoc T1.ide_w vseq_implies_hpar(1)
by auto
also have "... = composite_cell w \<theta> \<cdot> \<nu>"
proof -
have "(((r \<star> inv \<phi>) \<star> w) \<cdot> ((r \<star> \<phi>) \<star> w)) \<cdot> (\<rho> \<star> w) = \<rho> \<star> w"
proof -
have "\<guillemotleft>r \<star> \<phi> : r \<star> f \<Rightarrow> r \<star> f'\<guillemotright>"
using assms(1) by (intro hcomp_in_vhom, auto)
moreover have "\<guillemotleft>r \<star> inv \<phi> : r \<star> f' \<Rightarrow> r \<star> f\<guillemotright>"
using assms by (intro hcomp_in_vhom, auto)
ultimately show ?thesis
using comp_cod_arr
by (metis T1.w_in_hom(2) tab_simps(1) tab_simps(5) assms(1-2) comp_inv_arr'
in_homE leg0_simps(2) interchange base_in_hom(2) seqI')
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<omega>"
using w\<theta>\<nu> by simp
finally show ?thesis by simp
qed
qed
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T'.composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
by blast
qed
show "\<And>u w w' \<theta> \<theta>' \<beta>. \<lbrakk> ide w; ide w'; \<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f' \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>;
T'.composite_cell w \<theta> = T'.composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume w: "ide w" and w': "ide w'"
and \<theta>: "\<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> u\<guillemotright>" and \<theta>': "\<guillemotleft>\<theta>' : f' \<star> w' \<Rightarrow> u\<guillemotright>"
and \<beta>: "\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
and eq: "T'.composite_cell w \<theta> = T'.composite_cell w' \<theta>' \<cdot> \<beta>"
interpret uw\<theta>w'\<theta>'\<beta> V H \<a> \<i> src trg r \<open>(r \<star> \<phi>) \<cdot> \<rho>\<close> f' g u w \<theta> w' \<theta>' \<beta>
using w w' \<theta> \<theta>' \<beta> eq comp_assoc by (unfold_locales, auto)
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>)"
proof -
have \<phi>_w: "\<guillemotleft>\<phi> \<star> w : f \<star> w \<Rightarrow> f' \<star> w\<guillemotright>"
using assms(1) by (intro hcomp_in_vhom, auto)
have \<phi>_w': "\<guillemotleft>\<phi> \<star> w' : f \<star> w' \<Rightarrow> f' \<star> w'\<guillemotright>"
using assms(1) by (intro hcomp_in_vhom, auto)
have "\<guillemotleft>\<theta> \<cdot> (\<phi> \<star> w) : f \<star> w \<Rightarrow> u\<guillemotright>"
using \<theta> assms(1) by fastforce
moreover have "\<guillemotleft>\<theta>' \<cdot> (\<phi> \<star> w') : f \<star> w' \<Rightarrow> u\<guillemotright>"
using \<theta>' assms(1) by fastforce
moreover have "composite_cell w (\<theta> \<cdot> (\<phi> \<star> w)) = composite_cell w' (\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> \<beta>"
proof -
have "composite_cell w (\<theta> \<cdot> (\<phi> \<star> w)) =
(r \<star> \<theta>) \<cdot> ((r \<star> \<phi> \<star> w) \<cdot> \<a>[r, f, w]) \<cdot> (\<rho> \<star> w)"
using assms(2) \<phi>_w \<theta> whisker_left comp_assoc by auto
also have "... = (r \<star> \<theta>) \<cdot> \<a>[r, f', w] \<cdot> ((r \<star> \<phi>) \<star> w) \<cdot> (\<rho> \<star> w)"
using assms(1) assoc_naturality [of r \<phi> w] comp_assoc
by (metis \<phi>_w T'.tab_simps(1) base_simps(3) base_simps(4) hseq_char
in_homE seqE uw\<theta>.w_simps(5) uw\<theta>.w_simps(6))
also have "... = T'.composite_cell w \<theta>"
using assms(2) w whisker_right [of w] by simp
also have "... = T'.composite_cell w' \<theta>' \<cdot> \<beta>"
using eq by simp
also have "... = (r \<star> \<theta>') \<cdot> (\<a>[r, f', w'] \<cdot> ((r \<star> \<phi>) \<star> w')) \<cdot> (\<rho> \<star> w') \<cdot> \<beta>"
using assms(2) w' whisker_right [of w'] comp_assoc by simp
also have "... = ((r \<star> \<theta>') \<cdot> (r \<star> \<phi> \<star> w')) \<cdot> \<a>[r, f, w'] \<cdot> (\<rho> \<star> w') \<cdot> \<beta>"
using assms(1) assoc_naturality [of r \<phi> w'] comp_assoc
by (metis \<phi>_w' T'.tab_simps(1) base_simps(3) base_simps(4) hseqE in_homE seqE
uw'\<theta>'.w_simps(5) uw'\<theta>'.w_simps(6))
also have "... = composite_cell w' (\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> \<beta>"
using assms(2) whisker_left [of r] \<open>\<guillemotleft>\<theta>' \<cdot> (\<phi> \<star> w') : f \<star> w' \<Rightarrow> u\<guillemotright>\<close> comp_assoc
by auto
finally show ?thesis by simp
qed
ultimately have *: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and>
\<theta> \<cdot> (\<phi> \<star> w) = (\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> (f \<star> \<gamma>)"
using w w' \<beta> T2 by auto
show ?thesis
proof -
have **: "\<And>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<Longrightarrow> \<theta>' \<cdot> (\<phi> \<star> w') \<cdot> (f \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w) = \<theta>' \<cdot> (f' \<star> \<gamma>)"
proof -
fix \<gamma>
assume \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
have "\<theta>' \<cdot> (\<phi> \<star> w') \<cdot> (f \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w) = \<theta>' \<cdot> (\<phi> \<star> w') \<cdot> (f \<cdot> inv \<phi> \<star> \<gamma> \<cdot> w)"
using \<gamma> assms(1-2) interchange
by (metis arr_inv cod_inv in_homE leg0_simps(2) leg0_simps(4) uw\<theta>.w_in_hom(2)
seqI)
also have "... = \<theta>' \<cdot> (\<phi> \<cdot> f \<cdot> inv \<phi> \<star> w' \<cdot> \<gamma> \<cdot> w)"
using assms(1-2) interchange
by (metis \<gamma> arr_inv cod_inv comp_arr_dom comp_cod_arr in_homE seqI)
also have "... = \<theta>' \<cdot> (f' \<star> \<gamma>)"
proof -
have "\<phi> \<cdot> f \<cdot> inv \<phi> = f'"
using assms(1-2) comp_cod_arr comp_arr_inv' by auto
moreover have "w' \<cdot> \<gamma> \<cdot> w = \<gamma>"
using \<gamma> comp_arr_dom comp_cod_arr by auto
ultimately show ?thesis by simp
qed
finally show "\<theta>' \<cdot> (\<phi> \<star> w') \<cdot> (f \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w) = \<theta>' \<cdot> (f' \<star> \<gamma>)" by simp
qed
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and>
\<theta> \<cdot> (\<phi> \<star> w) = (\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> (f \<star> \<gamma>)"
using * by blast
have "\<theta> = \<theta>' \<cdot> (\<phi> \<star> w') \<cdot> (f \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w)"
proof -
have "seq (\<theta>' \<cdot> (\<phi> \<star> w')) (f \<star> \<gamma>)"
using assms(2) \<phi>_w \<phi>_w' \<gamma> \<beta> \<theta>
apply (intro seqI)
apply auto
by (metis seqE seqI')
thus ?thesis
using assms \<phi>_w \<gamma> comp_assoc invert_side_of_triangle(2) iso_hcomp
by (metis hcomp_in_vhomE ide_is_iso inv_hcomp inv_ide w)
qed
hence "\<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>)"
using \<gamma> ** by simp
hence "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>)"
using \<gamma> by auto
moreover have "\<And>\<gamma> \<gamma>'. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>) \<and>
\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>')
\<Longrightarrow> \<gamma> = \<gamma>'"
proof -
fix \<gamma> \<gamma>'
assume A: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>) \<and>
\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>')"
have "\<theta> \<cdot> (\<phi> \<star> w) = (\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> (f \<star> \<gamma>)"
proof -
have "\<theta> = ((\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> (f \<star> \<gamma>)) \<cdot> (inv \<phi> \<star> w)"
using A ** comp_assoc by simp
thus ?thesis
using assms(1-2) A iso_inv_iso
by (metis comp_arr_dom comp_cod_arr in_homE comp_assoc interchange)
qed
moreover have "\<theta> \<cdot> (\<phi> \<star> w) = (\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> (f \<star> \<gamma>')"
proof -
have "\<theta> = ((\<theta>' \<cdot> (\<phi> \<star> w')) \<cdot> (f \<star> \<gamma>')) \<cdot> (inv \<phi> \<star> w)"
using A ** comp_assoc by auto
thus ?thesis
using assms(1-2) A iso_inv_iso
by (metis comp_arr_dom comp_cod_arr in_homE comp_assoc interchange)
qed
ultimately show "\<gamma> = \<gamma>'"
using A * by blast
qed
ultimately show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f' \<star> \<gamma>)"
by metis
qed
qed
qed
qed
qed
text \<open>
Similarly, tabulations are preserved under composition of an isomorphism with
the ``output leg''.
\<close>
lemma preserved_by_output_iso:
assumes "\<guillemotleft>\<phi> : g' \<Rightarrow> g\<guillemotright>" and "iso \<phi>"
shows "tabulation V H \<a> \<i> src trg r (\<rho> \<cdot> \<phi>) f g'"
proof -
have \<tau>\<phi>: "\<guillemotleft>\<rho> \<cdot> \<phi> : g' \<Rightarrow> r \<star> f\<guillemotright>"
using assms by auto
interpret T': tabulation_data V H \<a> \<i> src trg r \<open>\<rho> \<cdot> \<phi>\<close> f g'
using assms(2) \<tau>\<phi> by (unfold_locales, auto)
have \<phi>_in_hhom: "\<guillemotleft>\<phi> : src f \<rightarrow> trg r\<guillemotright>"
using assms src_cod [of \<phi>] trg_cod [of \<phi>]
by (elim in_homE, simp)
show ?thesis
proof
fix u \<omega>
assume u: "ide u" and \<omega>: "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
show "\<exists>w \<theta> \<nu>'. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu>' : dom \<omega> \<Rightarrow> g' \<star> w\<guillemotright> \<and> iso \<nu>' \<and>
T'.composite_cell w \<theta> \<cdot> \<nu>' = \<omega>"
proof -
obtain w \<theta> \<nu> where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and>
iso \<nu> \<and> composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
using u \<omega> T1 [of u \<omega>] by auto
interpret uw\<theta>\<omega>\<nu>: uw\<theta>\<omega>\<nu> V H \<a> \<i> src trg r \<rho> f g u w \<theta> \<omega> \<nu>
using w\<theta>\<nu> comp_assoc by (unfold_locales, auto)
let ?\<nu>' = "(inv \<phi> \<star> w) \<cdot> \<nu>"
have \<nu>': "\<guillemotleft>?\<nu>' : dom \<omega> \<Rightarrow> g' \<star> w\<guillemotright>"
using assms \<phi>_in_hhom uw\<theta>\<omega>\<nu>.\<nu>_in_hom
by (intro comp_in_homI, auto)
moreover have "iso ?\<nu>'"
using assms \<nu>' w\<theta>\<nu> \<phi>_in_hhom
by (intro iso_hcomp isos_compose) auto
moreover have "T'.composite_cell w \<theta> \<cdot> ?\<nu>' = \<omega>"
proof -
have "composite_cell w \<theta> \<cdot> ((\<phi> \<star> w) \<cdot> ?\<nu>') = \<omega>"
proof -
have "(\<phi> \<star> w) \<cdot> ?\<nu>' = \<nu>"
using assms \<nu>' \<phi>_in_hhom whisker_right comp_cod_arr comp_assoc
by (metis comp_arr_inv' in_homE leg1_simps(2) uw\<theta>\<omega>\<nu>.uw\<theta>\<omega>\<nu>)
thus ?thesis
using w\<theta>\<nu> by simp
qed
moreover have "(\<rho> \<cdot> \<phi> \<star> w) \<cdot> ?\<nu>' = (\<rho> \<star> w) \<cdot> ((\<phi> \<star> w) \<cdot> ?\<nu>')"
using assms \<phi>_in_hhom whisker_right comp_assoc by simp
ultimately show ?thesis
using comp_assoc by simp
qed
ultimately show ?thesis
using w\<theta>\<nu> by blast
qed
next
fix u w w' \<theta> \<theta>' \<beta>'
assume w: "ide w" and w': "ide w'"
and \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>" and \<theta>': "\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>"
and \<beta>': "\<guillemotleft>\<beta>' : g' \<star> w \<Rightarrow> g' \<star> w'\<guillemotright>"
and eq': "T'.composite_cell w \<theta> = T'.composite_cell w' \<theta>' \<cdot> \<beta>'"
interpret uw\<theta>w'\<theta>'\<beta>: uw\<theta>w'\<theta>'\<beta> V H \<a> \<i> src trg r \<open>\<rho> \<cdot> \<phi>\<close> f g' u w \<theta> w' \<theta>' \<beta>'
using assms w w' \<theta> \<theta>' \<beta>' eq' comp_assoc by (unfold_locales, auto)
let ?\<beta> = "(\<phi> \<star> w') \<cdot> \<beta>' \<cdot> (inv \<phi> \<star> w)"
have \<beta>: "\<guillemotleft>?\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
using assms \<phi>_in_hhom \<beta>'
by (intro comp_in_homI hcomp_in_vhom, auto)
have eq: "composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> ((\<phi> \<star> w') \<cdot> \<beta>' \<cdot> (inv \<phi> \<star> w))"
proof -
have "composite_cell w \<theta> = (r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> ((\<rho> \<star> w) \<cdot> (\<phi> \<star> w)) \<cdot> (inv \<phi> \<star> w)"
proof -
have "\<rho> \<star> w = (\<rho> \<star> w) \<cdot> (\<phi> \<star> w) \<cdot> (inv \<phi> \<star> w)"
using assms w \<phi>_in_hhom whisker_right comp_arr_dom comp_arr_inv'
by (metis tab_simps(1) tab_simps(4) in_homE leg1_simps(2))
thus ?thesis
using comp_assoc by simp
qed
also have "... = T'.composite_cell w \<theta> \<cdot> (inv \<phi> \<star> w)"
using assms \<phi>_in_hhom whisker_right comp_assoc by simp
also have "... = T'.composite_cell w' \<theta>' \<cdot> (\<beta>' \<cdot> (inv \<phi> \<star> w))"
using eq' comp_assoc by simp
also have "... = composite_cell w' \<theta>' \<cdot> ((\<phi> \<star> w') \<cdot> \<beta>' \<cdot> (inv \<phi> \<star> w))"
using assms \<phi>_in_hhom whisker_right comp_assoc by simp
finally show ?thesis by simp
qed
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
proof -
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> ?\<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
using assms w w' \<theta> \<theta>' \<beta> eq \<phi>_in_hhom T2 [of w w' \<theta> u \<theta>' ?\<beta>] by auto
have "\<beta>' = g' \<star> \<gamma>"
proof -
have "g \<star> \<gamma> = (\<phi> \<star> w') \<cdot> \<beta>' \<cdot> (inv \<phi> \<star> w)"
using \<gamma> by simp
hence "(inv \<phi> \<star> w') \<cdot> (g \<star> \<gamma>) = \<beta>' \<cdot> (inv \<phi> \<star> w)"
using assms w' \<beta> \<phi>_in_hhom invert_side_of_triangle arrI iso_hcomp
hseqE ide_is_iso inv_hcomp inv_ide seqE
by metis
hence "\<beta>' = (inv \<phi> \<star> w') \<cdot> (g \<star> \<gamma>) \<cdot> (\<phi> \<star> w)"
using assms w \<beta> \<phi>_in_hhom invert_side_of_triangle comp_assoc seqE
by (metis comp_arr_dom in_homE local.uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(4) whisker_right)
also have "... = (inv \<phi> \<star> w') \<cdot> (\<phi> \<star> \<gamma>)"
using assms \<phi>_in_hhom \<gamma> interchange comp_arr_dom comp_cod_arr
by (metis in_homE)
also have "... = g' \<star> \<gamma>"
using assms \<phi>_in_hhom \<gamma> interchange comp_inv_arr inv_is_inverse comp_cod_arr
by (metis arr_dom calculation in_homE)
finally show ?thesis by simp
qed
hence "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
using \<beta> \<gamma> by auto
moreover have "\<And>\<gamma> \<gamma>'. \<lbrakk> \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>);
\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>') \<rbrakk> \<Longrightarrow> \<gamma> = \<gamma>'"
proof -
have *: "\<And>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<Longrightarrow> (\<phi> \<star> w') \<cdot> (g' \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w) = g \<star> \<gamma>"
proof -
fix \<gamma>
assume \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
have "(\<phi> \<star> w') \<cdot> (g' \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w) = (\<phi> \<star> w') \<cdot> (inv \<phi> \<star> \<gamma>)"
using assms \<phi>_in_hhom \<gamma> interchange comp_arr_dom comp_cod_arr
by (metis arr_dom comp_inv_arr' in_homE invert_side_of_triangle(2))
also have "... = g \<star> \<gamma>"
using assms \<phi>_in_hhom interchange comp_arr_inv inv_is_inverse comp_cod_arr
by (metis \<gamma> comp_arr_inv' in_homE leg1_simps(2))
finally show "(\<phi> \<star> w') \<cdot> (g' \<star> \<gamma>) \<cdot> (inv \<phi> \<star> w) = g \<star> \<gamma>" by blast
qed
fix \<gamma> \<gamma>'
assume \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
and \<gamma>': "\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>')"
show "\<gamma> = \<gamma>'"
using w w' \<theta> \<theta>' \<beta> \<gamma> \<gamma>' eq * T2 by metis
qed
ultimately show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta>' = g' \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)" by blast
qed
qed
qed
text \<open>
Finally, tabulations are preserved by composition with an isomorphism on the ``base''.
\<close>
lemma is_preserved_by_base_iso:
assumes "\<guillemotleft>\<phi> : r \<Rightarrow> r'\<guillemotright>" and "iso \<phi>"
shows "tabulation V H \<a> \<i> src trg r' ((\<phi> \<star> f) \<cdot> \<rho>) f g"
proof -
have \<phi>f: "\<guillemotleft>\<phi> \<star> f : r \<star> f \<Rightarrow> r' \<star> f\<guillemotright>"
using assms ide_leg0 by auto
interpret T: tabulation_data V H \<a> \<i> src trg r' \<open>(\<phi> \<star> f) \<cdot> \<rho>\<close> f
proof
show ide_r': "ide r'" using assms by auto
show "ide f" using ide_leg0 by auto
show "\<guillemotleft>(\<phi> \<star> f) \<cdot> \<rho> : g \<Rightarrow> r' \<star> f\<guillemotright>"
using tab_in_hom \<phi>f by force
qed
show ?thesis
proof
have *: "\<And>u v w \<theta> \<nu>. \<lbrakk> ide u; ide v; ide w; \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<nu> : v \<Rightarrow> g \<star> w\<guillemotright> \<rbrakk> \<Longrightarrow>
((\<phi> \<star> u) \<cdot> (r \<star> \<theta>)) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> =
T.composite_cell w \<theta> \<cdot> \<nu>"
proof -
fix u v w \<theta> \<nu>
assume u: "ide u" and v: "ide v" and w: "ide w"
and \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>" and \<nu>: "\<guillemotleft>\<nu> : v \<Rightarrow> g \<star> w\<guillemotright>"
have fw: "hseq f w"
using \<theta> ide_dom [of \<theta>] by fastforce
have r\<theta>: "hseq r \<theta>"
using \<theta> ide_base ide_dom [of \<theta>] trg_dom [of \<theta>]
using arrI fw vconn_implies_hpar(2) by auto
have "((\<phi> \<star> u) \<cdot> (r \<star> \<theta>)) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> =
((r' \<star> \<theta>) \<cdot> (\<phi> \<star> f \<star> w)) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu>"
using assms u w ide_base ide_leg0 \<theta> interchange comp_arr_dom comp_cod_arr
by (metis r\<theta> hseq_char in_homE)
also have "... = (r' \<star> \<theta>) \<cdot> ((\<phi> \<star> f \<star> w) \<cdot> \<a>[r, f, w]) \<cdot> (\<rho> \<star> w) \<cdot> \<nu>"
using comp_assoc by simp
also have "... = (r' \<star> \<theta>) \<cdot> \<a>[r', f, w] \<cdot> (((\<phi> \<star> f) \<star> w) \<cdot> (\<rho> \<star> w)) \<cdot> \<nu>"
proof -
have "(\<phi> \<star> f \<star> w) \<cdot> \<a>[r, f, w] = \<a>[r', f, w] \<cdot> ((\<phi> \<star> f) \<star> w)"
using assms ide_leg0 w assoc_naturality [of \<phi> f w] fw by fastforce
thus ?thesis
using comp_assoc by simp
qed
also have "... = T.composite_cell w \<theta> \<cdot> \<nu>"
using assms ide_leg0 whisker_right fw T.tab_in_hom arrI w comp_assoc by auto
finally show "((\<phi> \<star> u) \<cdot> (r \<star> \<theta>)) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> = T.composite_cell w \<theta> \<cdot> \<nu>"
by simp
qed
show "\<And>u \<omega>'. \<lbrakk> ide u; \<guillemotleft>\<omega>' : dom \<omega>' \<Rightarrow> r' \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega>' \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<cdot> \<nu> = \<omega>'"
proof -
fix u v \<omega>'
assume u: "ide u" and \<omega>': "\<guillemotleft>\<omega>' : v \<Rightarrow> r' \<star> u\<guillemotright>"
have \<omega>: "\<guillemotleft>(inv \<phi> \<star> u) \<cdot> \<omega>' : v \<Rightarrow> r \<star> u\<guillemotright>"
proof
show "\<guillemotleft>\<omega>' : v \<Rightarrow> r' \<star> u\<guillemotright>" by fact
show "\<guillemotleft>inv \<phi> \<star> u : r' \<star> u \<Rightarrow> r \<star> u\<guillemotright>"
proof -
have "ide (r' \<star> u)"
using \<omega>' ide_cod by fastforce
hence "hseq r' u" by simp
thus ?thesis
using assms u by auto
qed
qed
have \<phi>u: "hseq \<phi> u"
using assms \<omega> hseqI
by (metis arrI ide_is_iso iso_hcomp iso_is_arr seqE seq_if_composable
src_inv u)
obtain w \<theta> \<nu> where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = (inv \<phi> \<star> u) \<cdot> \<omega>'"
using u \<omega> T1 [of u "(inv \<phi> \<star> u) \<cdot> \<omega>'"] \<phi>f in_homE seqI' by auto
interpret uw\<theta>\<omega>\<nu> V H \<a> \<i> src trg r \<rho> f g u w \<theta> \<open>(inv \<phi> \<star> u) \<cdot> \<omega>'\<close> \<nu>
using w\<theta>\<nu> \<omega> comp_assoc by (unfold_locales, auto)
have "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<cdot> \<nu> = \<omega>'"
proof -
have "\<omega>' = ((\<phi> \<star> u) \<cdot> (r \<star> \<theta>)) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu>"
proof -
have "seq (r \<star> \<theta>) (\<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu>)" by fastforce
moreover have "iso (inv \<phi> \<star> u)"
using assms u \<phi>u by auto
moreover have "inv (inv \<phi> \<star> u) = \<phi> \<star> u"
using assms u \<phi>u by auto
ultimately show ?thesis
using invert_side_of_triangle(1) w\<theta>\<nu> comp_assoc by metis
qed
also have "... = T.composite_cell w \<theta> \<cdot> \<nu>"
using u w\<theta>\<nu> * [of u v w \<theta> \<nu>] by force
finally have "\<omega>' = T.composite_cell w \<theta> \<cdot> \<nu>" by simp
thus ?thesis
using w\<theta>\<nu> by simp
qed
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<cdot> \<nu> = \<omega>'"
by blast
qed
show "\<And>u w w' \<theta> \<theta>' \<beta>. \<lbrakk> ide w; ide w'; \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>;
T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume w: "ide w" and w': "ide w'"
and \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>" and \<theta>': "\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>"
and \<beta>: "\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
and eq': "T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<cdot> \<beta>"
interpret T: uw\<theta>w'\<theta>'\<beta> V H \<a> \<i> src trg r' \<open>(\<phi> \<star> f) \<cdot> \<rho>\<close> f g u w \<theta> w' \<theta>' \<beta>
using w w' \<theta> \<theta>' \<beta> eq' comp_assoc
by (unfold_locales, auto)
have eq: "composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta>"
proof -
have "(\<phi> \<star> u) \<cdot> composite_cell w \<theta> = (\<phi> \<star> u) \<cdot> composite_cell w' \<theta>' \<cdot> \<beta>"
proof -
have "(\<phi> \<star> u) \<cdot> composite_cell w \<theta> =
((\<phi> \<star> u) \<cdot> (r \<star> \<theta>)) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w) \<cdot> (g \<star> w)"
proof -
have "\<guillemotleft>\<rho> \<star> w : g \<star> w \<Rightarrow> (r \<star> f) \<star> w\<guillemotright>"
using w by auto
thus ?thesis
using comp_arr_dom comp_assoc by auto
qed
also have "... = T.composite_cell w \<theta> \<cdot> (g \<star> w)"
using * [of u "g \<star> w" w \<theta> "g \<star> w"] by fastforce
also have "... = T.composite_cell w \<theta>"
proof -
have "\<guillemotleft>(\<phi> \<star> f) \<cdot> \<rho> \<star> w : g \<star> w \<Rightarrow> (r' \<star> f) \<star> w\<guillemotright>"
using assms by fastforce
thus ?thesis
using comp_arr_dom comp_assoc by auto
qed
also have "... = T.composite_cell w' \<theta>' \<cdot> \<beta>"
using eq' by simp
also have "... = ((\<phi> \<star> u) \<cdot> (r \<star> \<theta>')) \<cdot> \<a>[r, f, w'] \<cdot> (\<rho> \<star> w') \<cdot> \<beta>"
using * [of u "g \<star> w" w' \<theta>' \<beta>] by fastforce
finally show ?thesis
using comp_assoc by simp
qed
moreover have "iso (\<phi> \<star> u)"
using assms by auto
moreover have "seq (\<phi> \<star> u) ((r \<star> \<theta>) \<cdot> \<a>[r, f, w] \<cdot> (\<rho> \<star> w))"
proof -
have "\<guillemotleft>\<phi> \<star> u : r \<star> u \<Rightarrow> r' \<star> u\<guillemotright>"
using assms by (intro hcomp_in_vhom, auto)
thus ?thesis
using composite_cell_in_hom [of w u \<theta>] by auto
qed
moreover have "seq (\<phi> \<star> u) (composite_cell w' \<theta>' \<cdot> \<beta>)"
using assms ide_leg0 w w' \<theta> \<theta>' \<beta> calculation(1) calculation(3) by auto
ultimately show ?thesis
using monoE section_is_mono iso_is_section by metis
qed
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> (f \<star> \<gamma>)"
using w w' \<theta> \<theta>' \<beta> eq T2 by simp
qed
qed
qed
end
subsection "Canonical Tabulations"
text \<open>
If the 1-cell \<open>g \<star> f\<^sup>*\<close> has any tabulation \<open>(f, \<rho>, g)\<close>, then it has the canonical
tabulation obtained as the adjoint transpose of (the identity on) \<open>g \<star> f\<^sup>*\<close>.
\<close>
context map_in_bicategory
begin
lemma canonical_tabulation:
assumes "ide g" and "src f = src g"
and "\<exists>\<rho>. tabulation V H \<a> \<i> src trg (g \<star> f\<^sup>*) \<rho> f g"
shows "tabulation V H \<a> \<i> src trg (g \<star> f\<^sup>*) (trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)) f g"
proof -
have 1: "ide (g \<star> f\<^sup>*)"
using assms(1-2) ide_right antipar by simp
obtain \<rho> where \<rho>: "tabulation V H \<a> \<i> src trg (g \<star> f\<^sup>*) \<rho> f g"
using assms(3) by auto
interpret \<rho>: tabulation V H \<a> \<i> src trg \<open>g \<star> f\<^sup>*\<close> \<rho> f g
using \<rho> by auto
let ?\<psi> = "trnr\<^sub>\<epsilon> (g \<star> f\<^sup>*) \<rho>"
have 3: "\<guillemotleft>?\<psi> : g \<star> f\<^sup>* \<Rightarrow> g \<star> f\<^sup>*\<guillemotright> \<and> iso ?\<psi>"
using \<rho>.yields_isomorphic_representation by blast
hence "tabulation (\<cdot>) (\<star>) \<a> \<i> src trg (g \<star> f\<^sup>*) ((inv ?\<psi> \<star> f) \<cdot> \<rho>) f g"
using \<rho>.is_preserved_by_base_iso [of "inv ?\<psi>" "g \<star> f\<^sup>*"] by simp
moreover have "(inv ?\<psi> \<star> f) \<cdot> \<rho> = trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)"
proof -
have "(inv ?\<psi> \<star> f) \<cdot> \<rho> = ((inv ?\<psi> \<star> f) \<cdot> (?\<psi> \<star> f)) \<cdot> trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)"
using \<rho>.\<rho>_in_terms_of_rep comp_assoc by simp
also have "... = ((g \<star> f\<^sup>*) \<star> f) \<cdot> trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)"
proof -
have "src (inv ?\<psi>) = trg f"
using 3 antipar
by (metis \<rho>.leg0_simps(3) \<rho>.base_in_hom(2) seqI' src_inv vseq_implies_hpar(1))
hence "(inv ?\<psi> \<star> f) \<cdot> (?\<psi> \<star> f) = (g \<star> f\<^sup>*) \<star> f"
using 3 whisker_right [of f "inv ?\<psi>" ?\<psi>] inv_is_inverse comp_inv_arr by auto
thus ?thesis
using comp_cod_arr by simp
qed
also have "... = trnr\<^sub>\<eta> g (g \<star> f\<^sup>*)"
proof -
have "src (g \<star> f\<^sup>*) = trg f" by simp
moreover have "ide g" by simp
ultimately have "\<guillemotleft>trnr\<^sub>\<eta> g (g \<star> f\<^sup>*) : g \<Rightarrow> (g \<star> f\<^sup>*) \<star> f\<guillemotright>"
using 1 adjoint_transpose_right(1) ide_in_hom antipar by blast
thus ?thesis
using comp_cod_arr by blast
qed
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
end
subsection "Uniqueness of Tabulations"
text \<open>
We now intend to show that a tabulation of \<open>r\<close> is ``unique up to equivalence'',
which is a property that any proper bicategorical limit should have.
What do we mean by this, exactly?
If we have two tabulations \<open>(f, \<rho>)\<close> and \<open>(f', \<rho>')\<close> of the same 1-cell \<open>r\<close>, then this
induces \<open>\<guillemotleft>w : src f' \<rightarrow> src f\<guillemotright>\<close>, \<open>\<guillemotleft>w' : src f \<rightarrow> src f'\<guillemotright>\<close>, \<open>\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> f'\<guillemotright>\<close>, and
\<open>\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> f'\<guillemotright>\<close>, such that \<open>\<rho>'\<close> is recovered up to isomorphism \<open>\<guillemotleft>\<nu> : g' \<Rightarrow> g \<star> w\<guillemotright>\<close>
from \<open>(w, \<theta>)\<close> by composition with \<open>\<rho>\<close> and \<open>\<rho>\<close> is recovered up to isomorphism
\<open>\<guillemotleft>\<nu>' : g \<Rightarrow> g' \<star> w'\<guillemotright>\<close> from \<open>(w', \<theta>')\<close> by composition with \<open>\<rho>'\<close>.
This means that we obtain isomorphisms \<open>\<guillemotleft>(\<nu>' \<star> w') \<cdot> \<nu> : g' \<Rightarrow> g' \<star> w' \<star> w\<guillemotright>\<close> and
\<open>\<guillemotleft>(\<nu> \<star> w') \<cdot> \<nu>' : g \<Rightarrow> g \<star> w \<star> w'\<guillemotright>\<close>.
These isomorphisms then induce, via \<open>T2\<close>, unique 2-cells from \<open>src f'\<close> to \<open>w' \<star> w\<close>
and from \<open>src f\<close> to \<open>w \<star> w'\<close>, which must be isomorphisms, thus showing \<open>w\<close> and \<open>w'\<close> are
equivalence maps.
\<close>
context tabulation
begin
text \<open>
We will need the following technical lemma.
\<close>
lemma apex_equivalence_lemma:
assumes "\<guillemotleft>\<rho>' : g' \<Rightarrow> r \<star> f'\<guillemotright>"
and "ide w \<and> \<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> f\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<Rightarrow> g' \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(r \<star> \<theta>) \<cdot> \<a>[r, f', w] \<cdot> (\<rho>' \<star> w) \<cdot> \<nu> = \<rho>"
and "ide w' \<and> \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>\<nu>' : g' \<Rightarrow> g \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
(r \<star> \<theta>') \<cdot> \<a>[r, f, w'] \<cdot> (\<rho> \<star> w') \<cdot> \<nu>' = \<rho>'"
shows "\<exists>\<phi>. \<guillemotleft>\<phi> : src f \<Rightarrow> w' \<star> w\<guillemotright> \<and> iso \<phi>"
proof -
interpret T': uw\<theta>\<omega>\<nu> V H \<a> \<i> src trg r \<rho> f g f' w' \<theta>' \<rho>' \<nu>'
using assms(1,3) apply unfold_locales by auto
interpret T: tabulation_data V H \<a> \<i> src trg r \<rho>' f' g'
using assms(1,2) apply unfold_locales by auto
interpret T: uw\<theta>\<omega>\<nu> V H \<a> \<i> src trg r \<rho>' f' g' f w \<theta> \<rho> \<nu>
using assms(1,2) apply unfold_locales by auto
(* These next simps are very important. *)
have dom_\<nu> [simp]: "dom \<nu> = dom \<rho>"
using assms(2) by auto
have dom_\<nu>' [simp]: "dom \<nu>' = dom \<rho>'"
using assms(3) by auto
let ?\<nu>'\<nu> = "\<a>[dom \<rho>, w', w] \<cdot> (\<nu>' \<star> w) \<cdot> \<nu>"
have \<nu>'\<nu>: "\<guillemotleft>?\<nu>'\<nu> : dom \<rho> \<Rightarrow> dom \<rho> \<star> w' \<star> w\<guillemotright>"
by fastforce
have "\<guillemotleft>\<nu> : src \<rho> \<rightarrow> trg r\<guillemotright>" by simp
let ?\<theta>\<theta>' = "\<theta> \<cdot> (\<theta>' \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[f, w', w]"
have \<theta>\<theta>': "\<guillemotleft>?\<theta>\<theta>' : f \<star> w' \<star> w \<Rightarrow> f\<guillemotright>"
by fastforce
have iso_\<nu>'\<nu>_r: "iso (?\<nu>'\<nu> \<cdot> \<r>[g])"
using iso_runit \<nu>'\<nu>
apply (intro isos_compose) by auto
have eq: "composite_cell (src f) \<r>[f] = composite_cell (w' \<star> w) ?\<theta>\<theta>' \<cdot> (?\<nu>'\<nu> \<cdot> \<r>[g])"
proof -
have "composite_cell (w' \<star> w) ?\<theta>\<theta>' \<cdot> (?\<nu>'\<nu> \<cdot> \<r>[g]) =
((r \<star> \<theta>) \<cdot> (r \<star> \<theta>' \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, w', w])) \<cdot>
\<a>[r, f, w' \<star> w] \<cdot> ((\<rho> \<star> w' \<star> w) \<cdot> \<a>[g, w', w]) \<cdot> (\<nu>' \<star> w) \<cdot> \<nu> \<cdot> \<r>[g]"
using whisker_left comp_assoc by simp
also have "... = ((r \<star> \<theta>) \<cdot> (r \<star> \<theta>' \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[f, w', w])) \<cdot>
\<a>[r, f, w' \<star> w] \<cdot> (\<a>[r \<star> f, w', w] \<cdot>
((\<rho> \<star> w') \<star> w)) \<cdot> (\<nu>' \<star> w) \<cdot> \<nu> \<cdot> \<r>[g]"
using assoc_naturality [of \<rho> w' w] by simp
also have "... = (r \<star> \<theta>) \<cdot> (r \<star> \<theta>' \<star> w) \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[f, w', w]) \<cdot> \<a>[r, f, w' \<star> w] \<cdot> \<a>[r \<star> f, w', w]) \<cdot>
((\<rho> \<star> w') \<star> w) \<cdot> (\<nu>' \<star> w) \<cdot> \<nu> \<cdot> \<r>[g]"
using comp_assoc by simp
also have "... = (r \<star> \<theta>) \<cdot> ((r \<star> \<theta>' \<star> w) \<cdot> \<a>[r, f \<star> w', w]) \<cdot>
(\<a>[r, f, w'] \<star> w) \<cdot>
((\<rho> \<star> w') \<star> w) \<cdot> (\<nu>' \<star> w) \<cdot> \<nu> \<cdot> \<r>[g]"
proof -
have "seq \<a>[r, f, w' \<star> w] \<a>[r \<star> f, w', w]" by simp
moreover have "inv (r \<star> \<a>[f, w', w]) = r \<star> \<a>\<^sup>-\<^sup>1[f, w', w]"
by simp
moreover have "(r \<star> \<a>[f, w', w]) \<cdot> \<a>[r, f \<star> w', w] \<cdot> (\<a>[r, f, w'] \<star> w) =
\<a>[r, f, w' \<star> w] \<cdot> \<a>[r \<star> f, w', w]"
using pentagon by simp
ultimately have "(r \<star> \<a>\<^sup>-\<^sup>1[f, w', w]) \<cdot> \<a>[r, f, w' \<star> w] \<cdot> \<a>[r \<star> f, w', w] =
\<a>[r, f \<star> w', w] \<cdot> (\<a>[r, f, w'] \<star> w)"
using iso_assoc [of f w' w] iso_hcomp
invert_side_of_triangle(1)
[of "\<a>[r, f, w' \<star> w] \<cdot> \<a>[r \<star> f, w', w]" "r \<star> \<a>[f, w', w]"
"\<a>[r, f \<star> w', w] \<cdot> (\<a>[r, f, w'] \<star> w)"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (r \<star> \<theta>) \<cdot> \<a>[r, f', w] \<cdot>
(((r \<star> \<theta>') \<star> w) \<cdot> (\<a>[r, f, w'] \<star> w) \<cdot> ((\<rho> \<star> w') \<star> w)) \<cdot>
(\<nu>' \<star> w) \<cdot> \<nu> \<cdot> \<r>[g]"
proof -
have "(r \<star> \<theta>' \<star> w) \<cdot> \<a>[r, f \<star> w', w] = \<a>[r, f', w] \<cdot> ((r \<star> \<theta>') \<star> w)"
using assoc_naturality [of r \<theta>' w] by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (r \<star> \<theta>) \<cdot> \<a>[r, f', w] \<cdot> (composite_cell w' \<theta>' \<star> w) \<cdot> (\<nu>' \<star> w) \<cdot> \<nu> \<cdot> \<r>[g]"
using whisker_right
by (metis T'.uw\<theta>\<omega> T'.w_in_hom(1) composite_cell_in_hom T'.\<theta>_simps(2) T'.ide_w
T.ide_w arrI seqE)
also have "... = (r \<star> \<theta>) \<cdot> \<a>[r, f', w] \<cdot> ((\<rho>' \<cdot> inv \<nu>' \<star> w) \<cdot> (\<nu>' \<star> w)) \<cdot> \<nu> \<cdot> \<r>[g]"
proof -
have "composite_cell w' \<theta>' = \<rho>' \<cdot> inv \<nu>'"
using assms invert_side_of_triangle(2) T.tab_simps(1) comp_assoc by presburger
thus ?thesis
using comp_assoc by simp
qed
also have "... = (T.composite_cell w \<theta> \<cdot> \<nu>) \<cdot> \<r>[g]"
using whisker_right [of w "\<rho>' \<cdot> inv \<nu>'" \<nu>'] dom_\<nu>' comp_assoc comp_inv_arr'
comp_arr_dom
by simp
also have "... = \<rho> \<cdot> \<r>[g]"
using assms(2) comp_assoc by simp
also have "... = composite_cell (src f) \<r>[f]"
using comp_assoc runit_hcomp runit_naturality [of \<rho>] by simp
finally show ?thesis by simp
qed
have eq': "(r \<star> \<r>[f]) \<cdot> \<a>[r, f, src f] \<cdot> (\<rho> \<star> src f) \<cdot> (inv (?\<nu>'\<nu> \<cdot> \<r>[g])) =
composite_cell (w' \<star> w) ?\<theta>\<theta>'"
proof -
have 1: "composite_cell (src f) \<r>[f] = (composite_cell (w' \<star> w) ?\<theta>\<theta>') \<cdot> ?\<nu>'\<nu> \<cdot> \<r>[g]"
using eq comp_assoc by simp
have "composite_cell (src f) \<r>[f] \<cdot> (inv (?\<nu>'\<nu> \<cdot> \<r>[g])) = composite_cell (w' \<star> w) ?\<theta>\<theta>'"
proof -
have "seq (r \<star> \<r>[f]) (\<a>[r, f, src f] \<cdot> (\<rho> \<star> src f))"
by fastforce
thus ?thesis
using iso_\<nu>'\<nu>_r 1 invert_side_of_triangle(2) by simp
qed
thus ?thesis
using comp_assoc by simp
qed
have \<nu>'\<nu>_r: "\<guillemotleft>?\<nu>'\<nu> \<cdot> \<r>[g] : g \<star> src f \<Rightarrow> g \<star> w' \<star> w\<guillemotright>"
by force
have inv_\<nu>'\<nu>_r: "\<guillemotleft>inv (?\<nu>'\<nu> \<cdot> \<r>[g]) : g \<star> w' \<star> w \<Rightarrow> g \<star> src f\<guillemotright>"
using \<nu>'\<nu> iso_\<nu>'\<nu>_r by auto
let ?P = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : src f \<Rightarrow> w' \<star> w\<guillemotright> \<and> ?\<nu>'\<nu> \<cdot> \<r>[g] = dom \<rho> \<star> \<gamma> \<and> \<r>[f] = ?\<theta>\<theta>' \<cdot> (f \<star> \<gamma>)"
let ?\<gamma> = "THE \<gamma>. ?P \<gamma>"
have "?P ?\<gamma>"
proof -
have "\<exists>!\<gamma>. ?P \<gamma>"
using \<nu>'\<nu>_r \<theta>\<theta>' eq T2 [of "src f" "w' \<star> w" "\<r>[f]" f ?\<theta>\<theta>' "?\<nu>'\<nu> \<cdot> \<r>[g]"] by simp
thus ?thesis
using the1_equality [of ?P] by blast
qed
hence \<gamma>: "\<guillemotleft>?\<gamma> : src f \<rightarrow> src f\<guillemotright> \<and> ?P ?\<gamma>"
using vconn_implies_hpar(1-2) by auto
let ?P' = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : w' \<star> w \<Rightarrow> src f\<guillemotright> \<and> inv (?\<nu>'\<nu> \<cdot> \<r>[g]) = g \<star> \<gamma> \<and> ?\<theta>\<theta>' = \<r>[f] \<cdot> (f \<star> \<gamma>)"
let ?\<gamma>' = "THE \<gamma>. ?P' \<gamma>"
have "?P' ?\<gamma>'"
proof -
have "\<exists>!\<gamma>. ?P' \<gamma>"
using inv_\<nu>'\<nu>_r \<theta>\<theta>' eq'
T2 [of "w' \<star> w" "src f" "\<theta> \<cdot> (\<theta>' \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[f, w', w]" f] comp_assoc
by simp
thus ?thesis
using the1_equality [of ?P'] by blast
qed
hence \<gamma>': "\<guillemotleft>?\<gamma>' : src f \<rightarrow> src f\<guillemotright> \<and> ?P' ?\<gamma>'"
using vconn_implies_hpar(1-2) by auto
have "inverse_arrows ?\<gamma> ?\<gamma>'"
proof
let ?Q = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : src f \<Rightarrow> src f\<guillemotright> \<and> dom \<rho> \<star> src f = g \<star> \<gamma> \<and> \<r>[f] = \<r>[f] \<cdot> (f \<star> \<gamma>)"
have "\<exists>!\<gamma>. ?Q \<gamma>"
proof -
have "ide (src f)" by simp
moreover have "\<guillemotleft>\<r>[f] : f \<star> src f \<Rightarrow> f\<guillemotright>" by simp
moreover have "\<guillemotleft>dom \<rho> \<star> src f : g \<star> src f \<Rightarrow> g \<star> src f\<guillemotright>" by auto
moreover have "(\<rho> \<star> src f) \<cdot> (dom \<rho> \<star> src f) = \<rho> \<star> src f"
proof -
have "(\<rho> \<star> src \<rho>) \<cdot> (dom \<rho> \<star> src (dom \<rho>)) = \<rho> \<star> src \<rho>"
using R.as_nat_trans.is_natural_1 arr_dom tab_simps(1) by presburger
thus ?thesis
by simp
qed
ultimately show ?thesis
using comp_arr_dom T2 [of "src f" "src f" "\<r>[f]" f "\<r>[f]" "dom \<rho> \<star> src f"]
comp_assoc
by metis
qed
moreover have "?Q (src f)"
using comp_arr_dom by auto
moreover have "?Q (?\<gamma>' \<cdot> ?\<gamma>)"
proof (intro conjI)
show "\<guillemotleft>?\<gamma>' \<cdot> ?\<gamma> : src f \<Rightarrow> src f\<guillemotright>"
using \<gamma> \<gamma>' by auto
show "dom \<rho> \<star> src f = g \<star> ?\<gamma>' \<cdot> ?\<gamma>"
proof -
have "g \<star> ?\<gamma>' \<cdot> ?\<gamma> = (g \<star> ?\<gamma>') \<cdot> (g \<star> ?\<gamma>)"
using \<gamma> \<gamma>' whisker_left by fastforce
also have "... = inv (?\<nu>'\<nu> \<cdot> \<r>[g]) \<cdot> (?\<nu>'\<nu> \<cdot> \<r>[g])"
using \<gamma> \<gamma>' by simp
also have "... = dom \<rho> \<star> src f"
using \<nu>'\<nu> iso_\<nu>'\<nu>_r comp_inv_arr inv_is_inverse by auto
finally show ?thesis by simp
qed
show "\<r>[f] = \<r>[f] \<cdot> (f \<star> ?\<gamma>' \<cdot> ?\<gamma>)"
proof -
have "\<r>[f] \<cdot> (f \<star> ?\<gamma>' \<cdot> ?\<gamma>) = \<r>[f] \<cdot> (f \<star> ?\<gamma>') \<cdot> (f \<star> ?\<gamma>)"
using \<gamma> \<gamma>' whisker_left by fastforce
also have "... = (\<r>[f] \<cdot> (f \<star> ?\<gamma>')) \<cdot> (f \<star> ?\<gamma>)"
using comp_assoc by simp
also have "... = \<r>[f]"
using \<gamma> \<gamma>' by simp
finally show ?thesis by simp
qed
qed
ultimately have "?\<gamma>' \<cdot> ?\<gamma> = src f" by blast
thus "ide (?\<gamma>' \<cdot> ?\<gamma>)" by simp
let ?Q' = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : w' \<star> w \<Rightarrow> w' \<star> w\<guillemotright> \<and> g \<star> w' \<star> w = g \<star> \<gamma> \<and> ?\<theta>\<theta>' = ?\<theta>\<theta>' \<cdot> (f \<star> \<gamma>)"
have "\<exists>!\<gamma>. ?Q' \<gamma>"
proof -
have "ide (w' \<star> w)" by simp
moreover have "\<guillemotleft>?\<theta>\<theta>' : f \<star> w' \<star> w \<Rightarrow> f\<guillemotright>"
using \<theta>\<theta>' by simp
moreover have "\<guillemotleft>g \<star> w' \<star> w : g \<star> w' \<star> w \<Rightarrow> g \<star> w' \<star> w\<guillemotright>"
by auto
moreover have
"composite_cell (w' \<star> w) ?\<theta>\<theta>' = composite_cell (w' \<star> w) ?\<theta>\<theta>' \<cdot> (g \<star> w' \<star> w)"
proof -
have "\<guillemotleft>\<rho> \<star> w' \<star> w : g \<star> w' \<star> w \<Rightarrow> (r \<star> f) \<star> w' \<star> w\<guillemotright>"
by (intro hcomp_in_vhom, auto)
hence "(\<rho> \<star> w' \<star> w) \<cdot> (g \<star> w' \<star> w) = \<rho> \<star> w' \<star> w"
using comp_arr_dom by auto
thus ?thesis
using comp_assoc by simp
qed
ultimately show ?thesis
using T2 by presburger
qed
moreover have "?Q' (w' \<star> w)"
using \<theta>\<theta>' comp_arr_dom by auto
moreover have "?Q' (?\<gamma> \<cdot> ?\<gamma>')"
proof (intro conjI)
show "\<guillemotleft>?\<gamma> \<cdot> ?\<gamma>' : w' \<star> w \<Rightarrow> w' \<star> w\<guillemotright>"
using \<gamma> \<gamma>' by auto
show "g \<star> w' \<star> w = g \<star> ?\<gamma> \<cdot> ?\<gamma>'"
proof -
have "g \<star> ?\<gamma> \<cdot> ?\<gamma>' = (g \<star> ?\<gamma>) \<cdot> (g \<star> ?\<gamma>')"
using \<gamma> \<gamma>' whisker_left by fastforce
also have "... = (?\<nu>'\<nu> \<cdot> \<r>[g]) \<cdot> inv (?\<nu>'\<nu> \<cdot> \<r>[g])"
using \<gamma> \<gamma>' by simp
also have "... = g \<star> w' \<star> w"
using \<nu>'\<nu> iso_\<nu>'\<nu>_r comp_arr_inv inv_is_inverse by auto
finally show ?thesis by simp
qed
show "?\<theta>\<theta>' = ?\<theta>\<theta>' \<cdot> (f \<star> ?\<gamma> \<cdot> ?\<gamma>')"
proof -
have "?\<theta>\<theta>' \<cdot> (f \<star> ?\<gamma> \<cdot> ?\<gamma>') = ?\<theta>\<theta>' \<cdot> (f \<star> ?\<gamma>) \<cdot> (f \<star> ?\<gamma>')"
using \<gamma> \<gamma>' whisker_left by fastforce
also have "... = (?\<theta>\<theta>' \<cdot> (f \<star> ?\<gamma>)) \<cdot> (f \<star> ?\<gamma>')"
using comp_assoc by simp
also have "... = ?\<theta>\<theta>'"
using \<gamma> \<gamma>' by simp
finally show ?thesis by simp
qed
qed
ultimately have "?\<gamma> \<cdot> ?\<gamma>' = w' \<star> w" by blast
thus "ide (?\<gamma> \<cdot> ?\<gamma>')" by simp
qed
hence "\<guillemotleft>?\<gamma> : src f \<Rightarrow> w' \<star> w\<guillemotright> \<and> iso ?\<gamma>"
using \<gamma> by auto
thus ?thesis by auto
qed
text \<open>
Now we can show that, given two tabulations of the same 1-cell,
there is an equivalence map between the apexes that extends to a transformation
of one tabulation into the other.
\<close>
lemma apex_unique_up_to_equivalence:
assumes "tabulation V H \<a> \<i> src trg r \<rho>' f' g'"
shows "\<exists>w w' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'.
equivalence_in_bicategory V H \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src f \<rightarrow> src f'\<guillemotright> \<and> \<guillemotleft>w' : src f' \<rightarrow> src f\<guillemotright> \<and>
\<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> f\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<Rightarrow> g' \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<rho> = (r \<star> \<theta>) \<cdot> \<a>[r, f', w] \<cdot> (\<rho>' \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>\<nu>' : g' \<Rightarrow> g \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
\<rho>' = (r \<star> \<theta>') \<cdot> \<a>[r, f, w'] \<cdot> (\<rho> \<star> w') \<cdot> \<nu>'"
proof -
interpret T': tabulation V H \<a> \<i> src trg r \<rho>' f' g'
using assms by auto
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : f' \<star> w \<Rightarrow> f\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<Rightarrow> g' \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<rho> = T'.composite_cell w \<theta> \<cdot> \<nu>"
using T'.T1 [of f \<rho>] ide_leg0 tab_in_hom by auto
obtain w' \<theta>' \<nu>'
where w'\<theta>'\<nu>': "ide w' \<and> \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>\<nu>' : g' \<Rightarrow> g \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
\<rho>' = composite_cell w' \<theta>' \<cdot> \<nu>'"
using T1 [of f' \<rho>'] T'.ide_leg0 T'.tab_in_hom by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : src f \<Rightarrow> w' \<star> w\<guillemotright> \<and> iso \<phi>"
using w\<theta>\<nu> w'\<theta>'\<nu>' apex_equivalence_lemma T'.tab_in_hom comp_assoc by metis
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : src f' \<Rightarrow> w \<star> w'\<guillemotright> \<and> iso \<psi>"
using w\<theta>\<nu> w'\<theta>'\<nu>' T'.apex_equivalence_lemma tab_in_hom comp_assoc by metis
have 1: "src f = src w"
using \<phi> src_dom [of \<phi>] hcomp_simps(1) [of w' w]
by (metis arr_cod in_homE leg0_simps(2) src_hcomp src_src vconn_implies_hpar(3))
have 2: "src f' = src w'"
using \<psi> src_dom [of \<psi>] hcomp_simps(1) [of w w']
by (metis T'.leg0_simps(2) arr_cod in_homE src_hcomp src_src vconn_implies_hpar(3))
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg w' w \<psi> \<open>inv \<phi>\<close>
using \<phi> \<psi> 1 2 w\<theta>\<nu> w'\<theta>'\<nu>' by unfold_locales auto
have "\<guillemotleft>w : src f \<rightarrow> src f'\<guillemotright>"
using \<psi> w\<theta>\<nu> 1 2 trg_cod hcomp_simps(2) E.antipar(1) by simp
moreover have "\<guillemotleft>w' : src f' \<rightarrow> src f\<guillemotright>"
using \<phi> w'\<theta>'\<nu>' 1 2 E.antipar(2) by simp
ultimately show ?thesis
using E.equivalence_in_bicategory_axioms w\<theta>\<nu> w'\<theta>'\<nu>' comp_assoc by metis
qed
end
subsection "`Tabulation' is Bicategorical"
text \<open>
In this section we show that ``tabulation'' is a truly bicategorical notion,
in the sense that tabulations are preserved and reflected by equivalence pseudofunctors.
The proofs given here is are elementary proofs from first principles.
It should also be possible to give a proof based on birepresentations,
but for this to actually save work it would first be necessary to carry out a general
development of birepresentations and bicategorical limits, and I have chosen not to
attempt this here.
\<close>
context equivalence_pseudofunctor
begin
lemma preserves_tabulation:
assumes "tabulation (\<cdot>\<^sub>C) (\<star>\<^sub>C) \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> f g"
shows "tabulation (\<cdot>\<^sub>D) (\<star>\<^sub>D) \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F r) (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>) (F f) (F g)"
proof -
let ?\<rho>' = "D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>"
interpret T: tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> f
using assms by auto
interpret T': tabulation_data V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F r\<close> ?\<rho>' \<open>F f\<close> \<open>F g\<close>
- using cmp_in_hom \<Phi>.components_are_iso C.VV.ide_char C.VV.arr_char
+ using cmp_in_hom \<Phi>.components_are_iso C.VV.ide_char\<^sub>S\<^sub>b\<^sub>C C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
apply unfold_locales
apply auto
by (intro D.comp_in_homI, auto)
interpret T': tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F r\<close> ?\<rho>' \<open>F f\<close> \<open>F g\<close>
text \<open>
How bad can it be to just show this directly from first principles?
It is worse than it at first seems, once you start filling in the details!
\<close>
proof
fix u' \<omega>'
assume u': "D.ide u'"
assume \<omega>': "\<guillemotleft>\<omega>' : D.dom \<omega>' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D u'\<guillemotright>"
show "\<exists>w' \<theta>' \<nu>'. D.ide w' \<and> \<guillemotleft>\<theta>' : F f \<star>\<^sub>D w' \<Rightarrow>\<^sub>D u'\<guillemotright> \<and>
\<guillemotleft>\<nu>' : D.dom \<omega>' \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w'\<guillemotright> \<and> D.iso \<nu>' \<and>
T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<nu>' = \<omega>'"
proof -
text \<open>
First, obtain \<open>\<omega>\<close> in \<open>C\<close> such that \<open>F \<omega>\<close> is related to \<open>\<omega>'\<close> by an equivalence in \<open>D\<close>.
\<close>
define v' where "v' = D.dom \<omega>'"
have v': "D.ide v'"
using assms v'_def D.ide_dom \<omega>' by blast
have \<omega>': "\<guillemotleft>\<omega>' : v' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D u'\<guillemotright>"
using v'_def \<omega>' by simp
define a' where "a' = src\<^sub>D \<omega>'"
have [simp]: "src\<^sub>D u' = a'"
using a'_def \<omega>'
by (metis D.arr_cod D.ide_char D.in_homE D.src.preserves_cod D.src_dom
D.src_hcomp v')
have [simp]: "trg\<^sub>D u' = src\<^sub>D (F r)"
using \<omega>'
by (metis D.cod_trg D.in_homE D.not_arr_null D.seq_if_composable D.trg.is_extensional
D.trg.preserves_arr D.trg.preserves_cod)
have [simp]: "src\<^sub>D v' = a'"
using v'_def \<omega>' a'_def by auto
have [simp]: "trg\<^sub>D v' = trg\<^sub>D (F r)"
using v'_def D.vconn_implies_hpar(4) \<omega>' u' by force
have [simp]: "src\<^sub>D \<omega>' = a'"
using \<omega>' a'_def by blast
have [simp]: "trg\<^sub>D \<omega>' = trg\<^sub>D (F r)"
using \<omega>' v'_def \<open>trg\<^sub>D v' = trg\<^sub>D (F r)\<close> by auto
obtain a where a: "C.obj a \<and> D.equivalent_objects (map\<^sub>0 a) a'"
using u' \<omega>' a'_def biessentially_surjective_on_objects D.obj_src by blast
obtain e' where e': "\<guillemotleft>e' : map\<^sub>0 a \<rightarrow>\<^sub>D a'\<guillemotright> \<and> D.equivalence_map e'"
using a D.equivalent_objects_def by auto
have u'_in_hhom: "\<guillemotleft>u' : a' \<rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C r)\<guillemotright>"
by (simp add: u')
hence 1: "\<guillemotleft>u' \<star>\<^sub>D e' : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C r)\<guillemotright>"
using e' by blast
have v'_in_hhom: "\<guillemotleft>v' : a' \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C r)\<guillemotright>"
by (simp add: v')
hence 2: "\<guillemotleft>v' \<star>\<^sub>D e' : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C r)\<guillemotright>"
using e' by blast
obtain d' \<eta>' \<epsilon>'
where d'\<eta>'\<epsilon>': "adjoint_equivalence_in_bicategory (\<cdot>\<^sub>D) (\<star>\<^sub>D) \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D e' d' \<eta>' \<epsilon>'"
using e' D.equivalence_map_extends_to_adjoint_equivalence by blast
interpret e': adjoint_equivalence_in_bicategory \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D e' d' \<eta>' \<epsilon>'
using d'\<eta>'\<epsilon>' by auto
interpret d': adjoint_equivalence_in_bicategory \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
d' e' "D.inv \<epsilon>'" "D.inv \<eta>'"
using e'.dual_adjoint_equivalence by simp
have [simp]: "src\<^sub>D e' = map\<^sub>0 a"
using e' by auto
have [simp]: "trg\<^sub>D e' = a'"
using e' by auto
have [simp]: "src\<^sub>D d' = a'"
by (simp add: e'.antipar(2))
have [simp]: "trg\<^sub>D d' = map\<^sub>0 a"
using e'.antipar by simp
obtain u where u: "\<guillemotleft>u : a \<rightarrow>\<^sub>C src\<^sub>C r\<guillemotright> \<and> C.ide u \<and> D.isomorphic (F u) (u' \<star>\<^sub>D e')"
using a e' u' 1 u'_in_hhom locally_essentially_surjective [of a "src\<^sub>C r" "u' \<star>\<^sub>D e'"]
C.obj_src D.equivalence_map_is_ide T.base_simps(2)
by blast
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : u' \<star>\<^sub>D e' \<Rightarrow>\<^sub>D F u\<guillemotright> \<and> D.iso \<phi>"
using u D.isomorphic_symmetric by blast
obtain v where v: "\<guillemotleft>v : a \<rightarrow>\<^sub>C trg\<^sub>C r\<guillemotright> \<and> C.ide v \<and> D.isomorphic (F v) (v' \<star>\<^sub>D e')"
using a e' v' v'_in_hhom locally_essentially_surjective [of a "trg\<^sub>C r" "v' \<star>\<^sub>D e'"]
C.obj_trg D.equivalence_map_is_ide T.base_simps(2)
by blast
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : F v \<Rightarrow>\<^sub>D v' \<star>\<^sub>D e'\<guillemotright> \<and> D.iso \<psi>"
using v by blast
have [simp]: "src\<^sub>C u = a" using u by auto
have [simp]: "trg\<^sub>C u = src\<^sub>C r" using u by auto
have [simp]: "src\<^sub>C v = a" using v by auto
have [simp]: "trg\<^sub>C v = trg\<^sub>C r" using v by auto
have [simp]: "src\<^sub>D \<phi> = map\<^sub>0 a"
using \<phi> by (metis "1" D.dom_src D.in_hhomE D.in_homE D.src.preserves_dom)
have [simp]: "trg\<^sub>D \<phi> = trg\<^sub>D u'"
using \<phi>
by (metis D.cod_trg D.hseqI D.in_homE D.isomorphic_implies_hpar(4)
D.trg.preserves_cod D.trg_hcomp e' u u'_in_hhom)
have [simp]: "src\<^sub>D \<psi> = map\<^sub>0 a"
using \<psi>
by (metis C.in_hhomE D.in_homE D.src_dom \<open>src\<^sub>D e' = map\<^sub>0 a\<close> preserves_src v)
have [simp]: "trg\<^sub>D \<psi> = trg\<^sub>D v'"
using \<psi>
by (metis "2" D.cod_trg D.in_hhomE D.in_homE D.trg.preserves_cod T.base_simps(2)
\<open>trg\<^sub>D v' = trg\<^sub>D (F r)\<close> preserves_trg)
define F\<omega> where "F\<omega> = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi>"
have F\<omega>: "\<guillemotleft>F\<omega> : F v \<Rightarrow>\<^sub>D F (r \<star>\<^sub>C u)\<guillemotright>"
proof (unfold F\<omega>_def, intro D.comp_in_homI)
show "\<guillemotleft>\<psi> : F v \<Rightarrow>\<^sub>D v' \<star>\<^sub>D e'\<guillemotright>"
using \<psi> by simp
show "\<guillemotleft>\<omega>' \<star>\<^sub>D e' : v' \<star>\<^sub>D e' \<Rightarrow>\<^sub>D (F r \<star>\<^sub>D u') \<star>\<^sub>D e'\<guillemotright>"
using e' \<omega>' D.equivalence_map_is_ide v'_in_hhom by blast
show "\<guillemotleft>\<a>\<^sub>D[F r, u', e'] : (F r \<star>\<^sub>D u') \<star>\<^sub>D e' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D u' \<star>\<^sub>D e'\<guillemotright>"
using e' u' D.equivalence_map_is_ide D.in_hhom_def u'_in_hhom by auto
show "\<guillemotleft>F r \<star>\<^sub>D \<phi> : F r \<star>\<^sub>D u' \<star>\<^sub>D e' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D F u\<guillemotright>"
using e' u' u \<phi>
by (metis C.in_hhomE D.hcomp_in_vhom D.isomorphic_implies_hpar(4)
T'.base_in_hom(2) T.base_simps(2) preserves_src preserves_trg)
show "\<guillemotleft>\<Phi> (r, u) : F r \<star>\<^sub>D F u \<Rightarrow>\<^sub>D F (r \<star>\<^sub>C u)\<guillemotright>"
using u cmp_in_hom(2) [of r u] by auto
qed
obtain \<omega> where \<omega>: "\<guillemotleft>\<omega> : v \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<guillemotright> \<and> F \<omega> = F\<omega>"
using u v \<omega>' \<phi> \<psi> F\<omega> locally_full [of v "r \<star>\<^sub>C u" F\<omega>]
by (metis C.ide_hcomp C.hseqI C.in_hhomE C.src_hcomp C.trg_hcomp
T.ide_base T.base_in_hom(1))
have [simp]: "src\<^sub>C \<omega> = src\<^sub>C u"
using \<omega>
by (metis C.hseqI C.in_homE C.src_cod C.src_hcomp T.base_in_hom(1) u)
have [simp]: "trg\<^sub>C \<omega> = trg\<^sub>C r"
using \<omega>
by (metis C.ide_char C.ide_trg C.in_homE C.trg.preserves_hom \<open>trg\<^sub>C v = trg\<^sub>C r\<close>)
text \<open>Apply \<open>T.T1\<close> to \<open>u\<close> and \<open>\<omega>\<close> to obtain \<open>w\<close>, \<open>\<theta>\<close>, \<open>\<nu>\<close>.\<close>
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "C.ide w \<and> \<guillemotleft>\<theta> : f \<star>\<^sub>C w \<Rightarrow>\<^sub>C u\<guillemotright> \<and> \<guillemotleft>\<nu> : C.dom \<omega> \<Rightarrow>\<^sub>C g \<star>\<^sub>C w\<guillemotright> \<and>
C.iso \<nu> \<and> T.composite_cell w \<theta> \<cdot>\<^sub>C \<nu> = \<omega>"
using u \<omega> T.T1 [of u \<omega>] by auto
text \<open>
Combining \<open>\<omega>\<close> and \<open>w\<theta>\<nu>\<close> yields the situation depicted in the diagram below.
In this as well as subsequent diagrams, canonical isomorphisms have been suppressed
in the interests of clarity.
$$
F (
\xy/67pt/
\xymatrix{
& {\scriptstyle{a}}
\xlowertwocell[ddddl]{}_{v}{^\nu}
\xuppertwocell[ddddr]{}^{u}{^\theta}
\ar@ {.>}[dd]^{w}
\\
\\
& \scriptstyle{{\rm src}~g \;=\;{\rm src}~f} \xtwocell[ddd]{}\omit{^\rho}
\ar[ddl] _{g}
\ar[ddr] ^{f}
\\
\\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll] ^{r}
\\
&
}
\endxy
)
\qquad = \qquad
\xy/67pt/
\xymatrix{
& {\scriptstyle{{\rm src}(F a)}}
\xlowertwocell[ddddl]{}^{<2>F v}{^{\psi}}
\xuppertwocell[ddddr]{}^{<2>F u}{^{\phi}}
\ar[dd] ^{e'}
\\
\\
& \scriptstyle{a'} \xtwocell[ddd]{}\omit{^{\omega'}}
\ar[ddl] _{v'}
\ar[ddr] ^{u'}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
$$
\<close>
have [simp]: "src\<^sub>C w = src\<^sub>C u"
by (metis C.arrI C.seqE C.src_hcomp C.src_vcomp C.vseq_implies_hpar(1)
\<omega> \<open>src\<^sub>C \<omega> = src\<^sub>C u\<close> w\<theta>\<nu>)
have [simp]: "trg\<^sub>C w = src\<^sub>C f"
by (metis C.arrI C.hseq_char C.seqE T.tab_simps(2) \<omega> w\<theta>\<nu>)
have [simp]: "src\<^sub>D (F u) = map\<^sub>0 a"
using e'.antipar(1) u by auto
have [simp]: "src\<^sub>D (F v) = map\<^sub>0 a"
using v e' e'.antipar by force
have [simp]: "src\<^sub>D (F w) = map\<^sub>0 a"
by (simp add: w\<theta>\<nu>)
have *: "F (T.composite_cell w \<theta> \<cdot>\<^sub>C \<nu>) =
\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu>"
text \<open>
$$
F (
\xy/67pt/
\xymatrix{
& {\scriptstyle{a}}
\xlowertwocell[ddddl]{}_{v}{^\nu}
\xuppertwocell[ddddr]{}^{u}{^\theta}
\ar[dd] ^{w}
\\
\\
& \scriptstyle{{\rm src}~g \;=\;{\rm src}~f} \xtwocell[ddd]{}\omit{^\rho}
\ar[ddl] _{g}
\ar[ddr] ^{f}
\\
\\
\scriptstyle{{\rm trg}~r} & & \scriptstyle{{\rm src}~r} \ar[ll] ^{r}
\\
&
}
\endxy
)
\qquad = \qquad
\xy/67pt/
\xymatrix{
& {\scriptstyle{{\rm src}(F a)}}
\xlowertwocell[ddddl]{}^{<2>F v}{^{F \nu}}
\xuppertwocell[ddddr]{}^{<2>F u}{^{F \theta}}
\ar[dd] ^{Fw}
\\
\\
& \scriptstyle{{\rm src}(F g) \;=\;{\rm src}(F f)} \xtwocell[ddd]{}\omit{^{F \rho}}
\ar[ddl] _{F g}
\ar[ddr] ^{F f}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
$$
\<close>
proof -
have "F (T.composite_cell w \<theta> \<cdot>\<^sub>C \<nu>) = F ((r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>C \<a>\<^sub>C[r, f, w] \<cdot>\<^sub>C (\<rho> \<star>\<^sub>C w) \<cdot>\<^sub>C \<nu>)"
using C.comp_assoc by simp
also have "... = F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D F \<a>\<^sub>C[r, f, w] \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w) \<cdot>\<^sub>D F \<nu>"
by (metis C.arr_dom_iff_arr C.comp_assoc C.in_homE C.seqE
as_nat_trans.preserves_comp_2 w\<theta>\<nu>)
also have "... =
F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D (\<Phi> (r, f \<star>\<^sub>C w) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w))) \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w) \<cdot>\<^sub>D F \<nu>"
using \<omega> w\<theta>\<nu> preserves_assoc [of r f w]
by (metis C.hseqE C.in_homE C.seqE T.tab_simps(2) T.ide_leg0 T.ide_base
T.leg0_simps(3))
also have "... =
((F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w))) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w))) \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w) \<cdot>\<^sub>D F \<nu>"
using D.comp_assoc by simp
also have "... =
\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w)) \<cdot>\<^sub>D F \<nu>"
proof -
have "(F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) =
(\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)))"
proof -
have "F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w) = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>)"
- using \<omega> \<Phi>.naturality [of "(r, \<theta>)"] FF_def w\<theta>\<nu> C.VV.arr_char
+ using \<omega> \<Phi>.naturality [of "(r, \<theta>)"] FF_def w\<theta>\<nu> C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
C.VV.dom_simp C.VV.cod_simp
apply simp
by (metis (no_types, lifting) C.hseqE C.in_homE C.seqE)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w))"
proof -
have "(F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) = F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)"
using \<omega> w\<theta>\<nu> D.whisker_right [of "F r" "F \<theta>" "\<Phi> (f, w)"]
by (metis C.hseqE C.in_homE C.seqE D.comp_ide_self D.interchange D.seqI'
T'.ide_base T'.base_in_hom(2) T.tab_simps(2) T.ide_leg0 cmp_in_hom(2)
preserves_hom)
thus ?thesis by simp
qed
finally have "(F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) =
\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w))"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))) \<cdot>\<^sub>D F \<nu>"
proof -
have "(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w) =
((D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w)) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
proof -
have "D.inv (\<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w) = (F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
proof -
have "src\<^sub>C (r \<star>\<^sub>C f) = trg\<^sub>C w"
using \<omega> w\<theta>\<nu>
by (metis C.arrI C.hseq_char C.seqE C.hcomp_simps(1) T.tab_simps(2)
T.leg0_simps(2) T.leg0_simps(3))
hence "D.seq (\<Phi> (r \<star>\<^sub>C f, w)) (F \<rho> \<star>\<^sub>D F w)"
- using \<omega> w\<theta>\<nu> cmp_in_hom(2) [of "r \<star>\<^sub>C f" w] C.VV.arr_char FF_def by auto
+ using \<omega> w\<theta>\<nu> cmp_in_hom(2) [of "r \<star>\<^sub>C f" w] C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C FF_def by auto
moreover have "\<Phi> (r \<star>\<^sub>C f, w) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w) = F (\<rho> \<star>\<^sub>C w) \<cdot>\<^sub>D \<Phi> (g, w)"
using \<omega> w\<theta>\<nu> \<Phi>.naturality [of "(\<rho>, w)"] cmp_components_are_iso FF_def
- C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
+ C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C C.VV.dom_simp C.VV.cod_simp
by simp
moreover have "D.iso (\<Phi> (r \<star>\<^sub>C f, w))"
using w\<theta>\<nu> cmp_components_are_iso
by (metis C.arrI C.ide_hcomp C.hseqE C.hseqI' C.seqE C.src_hcomp
T.tab_simps(2) T.ide_leg0 T.ide_base T.leg0_simps(2-3) \<omega>)
moreover have "D.iso (\<Phi> (g, w))"
using w\<theta>\<nu> cmp_components_are_iso
by (metis C.arrI C.hseqE C.seqE T.tab_simps(2) T.ide_leg1 T.leg1_simps(3) \<omega>)
ultimately show ?thesis
- using \<omega> w\<theta>\<nu> \<Phi>.naturality cmp_components_are_iso FF_def C.VV.arr_char
+ using \<omega> w\<theta>\<nu> \<Phi>.naturality cmp_components_are_iso FF_def C.VV.arr_char\<^sub>S\<^sub>b\<^sub>C
D.invert_opposite_sides_of_square
by presburger
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
using \<omega> w\<theta>\<nu> D.whisker_right cmp_components_are_iso cmp_in_hom D.comp_assoc
by auto
finally show ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis
using D.comp_assoc by simp
qed
text \<open>We can now define the \<open>w'\<close>, \<open>\<theta>'\<close>, and \<open>\<nu>'\<close> that we are required to exhibit.\<close>
define \<phi>' where "\<phi>' = e'.trnr\<^sub>\<epsilon> u' (D.inv \<phi>)"
have "\<phi>' = \<r>\<^sub>D[u'] \<cdot>\<^sub>D (u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[u', e', d'] \<cdot>\<^sub>D (D.inv \<phi> \<star>\<^sub>D d')"
unfolding \<phi>'_def e'.trnr\<^sub>\<epsilon>_def by simp
have \<phi>': "\<guillemotleft>\<phi>' : F u \<star>\<^sub>D d' \<Rightarrow>\<^sub>D u'\<guillemotright>"
using \<phi> \<phi>'_def u u' e'.adjoint_transpose_right(2) [of u' "F u"] by auto
have [simp]: "src\<^sub>D \<phi>' = src\<^sub>D u'"
using \<phi>' by fastforce
have [simp]: "trg\<^sub>D \<phi>' = trg\<^sub>D u'"
using \<phi>' by fastforce
define \<psi>' where "\<psi>' = d'.trnr\<^sub>\<eta> v' (D.inv \<psi>)"
have \<psi>'_eq: "\<psi>' = (D.inv \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d'] \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
unfolding \<psi>'_def d'.trnr\<^sub>\<eta>_def by simp
have \<psi>': "\<guillemotleft>\<psi>' : v' \<Rightarrow>\<^sub>D F v \<star>\<^sub>D d'\<guillemotright>"
using \<psi> \<psi>'_def v v' d'.adjoint_transpose_right(1) [of "F v" v'] by auto
have iso_\<psi>': "D.iso \<psi>'"
unfolding \<psi>'_def d'.trnr\<^sub>\<eta>_def
using \<psi> e'.counit_is_iso
by (metis D.arrI D.iso_hcomp D.hseq_char D.ide_is_iso D.iso_assoc'
D.iso_inv_iso D.iso_runit' D.isos_compose D.seqE \<psi>'_eq
\<psi>' d'.unit_simps(5) e'.antipar(1) e'.antipar(2) e'.ide_left e'.ide_right v')
have [simp]: "src\<^sub>D \<psi>' = src\<^sub>D v'"
using \<psi>' by fastforce
have [simp]: "trg\<^sub>D \<psi>' = trg\<^sub>D v'"
using \<psi>' by fastforce
define w' where "w' = F w \<star>\<^sub>D d'"
define \<theta>' where "\<theta>' = \<phi>' \<cdot>\<^sub>D (F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']"
define \<nu>' where "\<nu>' = \<a>\<^sub>D[F g, F w, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
have w': "D.ide w' \<and> \<guillemotleft>w' : src\<^sub>D u' \<rightarrow>\<^sub>D src\<^sub>D (F f)\<guillemotright>"
using w'_def \<omega> w\<theta>\<nu> by simp
have \<theta>': "\<guillemotleft>\<theta>' : F f \<star>\<^sub>D w' \<Rightarrow>\<^sub>D u'\<guillemotright>"
unfolding \<theta>'_def w'_def
using \<phi>' \<omega> w\<theta>\<nu> cmp_in_hom
apply (intro D.comp_in_homI D.hcomp_in_vhom)
apply auto
by (intro D.comp_in_homI D.hcomp_in_vhom, auto)
have \<nu>': "\<guillemotleft>\<nu>' : v' \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w'\<guillemotright>"
unfolding \<nu>'_def w'_def
using \<psi>' \<omega> w\<theta>\<nu> cmp_in_hom cmp_components_are_iso
apply (intro D.comp_in_homI)
apply auto
by (intro D.hcomp_in_vhom D.comp_in_homI, auto)
have iso_\<nu>': "D.iso \<nu>'"
using \<nu>'_def iso_\<psi>' cmp_in_hom D.isos_compose preserves_iso
by (metis (no_types, lifting) C.ideD(1) D.arrI D.iso_hcomp D.hseqE D.ide_is_iso
D.iso_assoc D.iso_inv_iso D.seqE T.ide_leg1 T.leg1_simps(3) cmp_components_are_iso
\<nu>' \<open>src\<^sub>D (F w) = map\<^sub>0 a\<close> \<open>src\<^sub>D e' = map\<^sub>0 a\<close> \<open>trg\<^sub>C w = src\<^sub>C f\<close> e'.antipar(1)
e'.ide_right preserves_ide preserves_src preserves_trg w\<theta>\<nu>)
have "T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<nu>' = \<omega>'"
text \<open>
$$
\xy/67pt/
\xymatrix{
&
\xlowertwocell[ddddddl]{\scriptstyle{a'}}<-13>^{<2>v'}{^{\psi'}}
\xuppertwocell[ddddddr]{}<13>^{<2>u'}{^{\phi'}}
\ar [dd] ^{d'}
\\
\\
& {\scriptstyle{{\rm src}(F g) \;=\;{\rm src}(F f)}}
\xlowertwocell[ddddl]{}^{<2>F v}{^{F \nu}}
\xuppertwocell[ddddr]{}^{<2>F u}{^{F \theta}}
\ar[dd] ^{Fw}
\\
\\
& \scriptstyle{a'} \xtwocell[ddd]{}\omit{^{F \rho}}
\ar[ddl] _{F g}
\ar[ddr] ^{F f}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
\qquad = \qquad
\xy/33pt/
\xymatrix{
& \scriptstyle{\scriptstyle{a'}} \xtwocell[ddd]{}\omit{^{\omega'}}
\ar[ddl] _{v'}
\ar[ddr] ^{u'}
\\
\\
\scriptstyle{{\rm trg}~(Fr)} & & \scriptstyle{{\rm src}~(Fr)} \ar[ll] ^{Fr}
\\
&
}
\endxy
$$
\<close>
proof -
have 1: "\<guillemotleft>T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<nu>' : v' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D u'\<guillemotright>"
using w' \<theta>' \<nu>' w\<theta>\<nu> T'.composite_cell_in_hom by blast
have "T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<nu>' =
(F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D
(F (T.composite_cell w \<theta> \<cdot>\<^sub>C \<nu>) \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<nu>' =
(F r \<star>\<^sub>D \<phi>' \<cdot>\<^sub>D (F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, w'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D w') \<cdot>\<^sub>D \<a>\<^sub>D[F g, F w, d'] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using \<theta>'_def \<nu>'_def D.comp_assoc by simp
also have
"... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D (F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F w, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using \<theta>' \<theta>'_def w'_def D.comp_assoc D.whisker_left by auto
also have
"... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D (F \<theta> \<star>\<^sub>D d') \<cdot>\<^sub>D (\<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F w, d']) \<cdot>\<^sub>D (D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using \<theta>' \<theta>'_def D.whisker_right cmp_in_hom D.comp_assoc by fastforce
also have
"... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D (F \<theta> \<star>\<^sub>D d') \<cdot>\<^sub>D (\<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<cdot>\<^sub>D
\<a>\<^sub>D[F r \<star>\<^sub>D F f, F w, d'] \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d') \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F g, F w, d'] =
\<a>\<^sub>D[F r \<star>\<^sub>D F f, F w, d'] \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d')"
using D.assoc_naturality [of "D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>" "F w" d']
cmp_in_hom cmp_components_are_iso
by (simp add: w\<theta>\<nu>)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<cdot>\<^sub>D \<a>\<^sub>D[F r \<star>\<^sub>D F f, F w, d']) \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d') \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using 1 D.whisker_left D.comp_assoc
by (metis D.arrI D.hseq_char D.seqE T'.ide_base calculation)
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<star>\<^sub>D d') \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f \<star>\<^sub>D F w, d']) \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, F w] \<star>\<^sub>D d') \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d') \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "D.seq \<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<a>\<^sub>D[F r \<star>\<^sub>D F f, F w, d']"
by (metis 1 D.arrI D.seqE calculation)
hence "(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F w, d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<cdot>\<^sub>D
\<a>\<^sub>D[F r \<star>\<^sub>D F f, F w, d'] =
\<a>\<^sub>D[F r, F f \<star>\<^sub>D F w, d'] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, F w] \<star>\<^sub>D d')"
using w\<theta>\<nu> D.pentagon
D.invert_side_of_triangle(1)
[of "\<a>\<^sub>D[F r, F f, F w \<star>\<^sub>D d'] \<cdot>\<^sub>D \<a>\<^sub>D[F r \<star>\<^sub>D F f, F w, d']"
"F r \<star>\<^sub>D \<a>\<^sub>D[F f, F w, d']"
"\<a>\<^sub>D[F r, F f \<star>\<^sub>D F w, d'] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, F w] \<star>\<^sub>D d')"]
by (simp add: w\<theta>\<nu>)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D ((F r \<star>\<^sub>D F \<theta> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F (f \<star>\<^sub>C w), d']) \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<Phi> (f, w)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, F w] \<star>\<^sub>D d') \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d') \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "(F r \<star>\<^sub>D \<Phi> (f, w) \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f \<star>\<^sub>D F w, d'] =
\<a>\<^sub>D[F r, F (f \<star>\<^sub>C w), d'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<Phi> (f, w)) \<star>\<^sub>D d')"
using 1 w\<theta>\<nu> D.assoc_naturality [of "F r" "\<Phi> (f, w)" d']
\<open>trg\<^sub>C w = src\<^sub>C f\<close> e'.ide_right
by (metis D.arrI D.hseq_char D.ide_char D.seqE T'.base_simps(3)
T'.base_simps(4) T'.leg0_simps(3) T.ide_leg0 cmp_simps(1-5) w'_def)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (((F r \<star>\<^sub>D F \<theta>) \<star>\<^sub>D d') \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<Phi> (f, w)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, F w] \<star>\<^sub>D d') \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d') \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d')) \<cdot>\<^sub>D \<psi>'"
proof -
have "src\<^sub>D (F r) = trg\<^sub>D (F \<theta>)"
using w\<theta>\<nu> by (metis C.arrI C.hseqE C.seqE \<omega> preserves_hseq)
moreover have "src\<^sub>D (F \<theta>) = trg\<^sub>D d'"
using w\<theta>\<nu> C.arrI C.vconn_implies_hpar(1) by auto
ultimately
have "(F r \<star>\<^sub>D F \<theta> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F (f \<star>\<^sub>C w), d'] =
\<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F \<theta>) \<star>\<^sub>D d')"
using w\<theta>\<nu> D.assoc_naturality [of "F r" "F \<theta>" d'] by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
(((F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w))) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "((F r \<star>\<^sub>D F \<theta>) \<star>\<^sub>D d') \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<Phi> (f, w)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, F w] \<star>\<^sub>D d') \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<star>\<^sub>D d') \<cdot>\<^sub>D
(D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') =
(F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu>
\<star>\<^sub>D d'"
proof -
have "\<guillemotleft>(F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> :
F v \<Rightarrow>\<^sub>D F r \<star>\<^sub>D F u\<guillemotright>"
using w\<theta>\<nu> \<omega> cmp_in_hom
apply (intro D.comp_in_homI)
apply auto
by (intro D.hcomp_in_vhom, auto)
hence "D.arr ((F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu>)"
by auto
thus ?thesis
using D.whisker_right by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using w\<theta>\<nu> D.whisker_left cmp_in_hom
by (metis D.seqI' T'.ide_base T.ide_leg0 \<open>trg\<^sub>C w = src\<^sub>C f\<close> preserves_hom)
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u) \<cdot>\<^sub>D
(F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w))) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "(D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u)) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) =
F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)"
proof -
have "(D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u)) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) =
(F r \<star>\<^sub>D F u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w))"
using u cmp_components_are_iso
by (simp add: D.comp_inv_arr')
also have "... = F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)"
using u \<omega> w\<theta>\<nu> cmp_in_hom \<open>trg\<^sub>C u = src\<^sub>C r\<close>
D.comp_cod_arr [of "F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)" "F r \<star>\<^sub>D F u"]
by (metis (full_types) "*" D.arrI D.cod_comp D.seqE F\<omega> T.ide_base
cmp_simps(4))
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using D.comp_assoc by simp
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D
(\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d' =
(D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D F \<nu> \<star>\<^sub>D d')"
using D.whisker_right cmp_in_hom cmp_components_are_iso
by (metis * D.arrI D.invert_side_of_triangle(1) F\<omega> T.ide_base \<omega>
\<open>trg\<^sub>C u = src\<^sub>C r\<close> e'.ide_right u w\<theta>\<nu>)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D
(F (T.composite_cell w \<theta> \<cdot>\<^sub>C \<nu>) \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using D.comp_assoc * by simp
finally show ?thesis by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D
(F \<omega> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using w\<theta>\<nu> by simp
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D
(\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D
\<psi>'"
using \<omega> F\<omega>_def by simp
text \<open>
$$
\xy/67pt/
\xymatrix{
& {\scriptstyle{a'}}
\xlowertwocell[ddddl]{}^{<2>F v}{^{\psi'}}
\xuppertwocell[ddddr]{}^{<2>F u}{^{\phi'}}
\ar@ {.}[dd] ^{d'}
\\
\\
& \scriptstyle{{\rm src}(F a)} \xtwocell[ddd]{}\omit{^{F \omega}}
\ar[ddl] _{F v}
\ar[ddr] ^{F u}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
\qquad = \qquad
\xy/67pt/
\xymatrix{
&
\xlowertwocell[ddddddl]{\scriptstyle{a'}}<-13>^{<2>v'}{^{\psi'}}
\xuppertwocell[ddddddr]{}<13>^{<2>u'}{^{\phi'}}
\ar@ {.}[dd] ^{d'}
\\
\\
& {\scriptstyle{{\rm src}(F a)}}
\xlowertwocell[ddddl]{}^{<2>F v}{^{\psi}}
\xuppertwocell[ddddr]{}^{<2>F u}{^{\phi}}
\ar@ {.}[dd] ^{e'}
\\
\\
& \scriptstyle{a'} \xtwocell[ddd]{}\omit{^{\omega'}}
\ar[ddl] _{v'}
\ar[ddr] ^{u'}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
$$
\<close>
also have "... = \<omega>'"
text \<open>
$$
\xy/67pt/
\xymatrix{
&
\xlowertwocell[ddddddl]{\scriptstyle{a'}}<-13>^{<2>v'}{^{\psi'}}
\xuppertwocell[ddddddr]{}<13>^{<2>u'}{^{\phi'}}
\ar[dd] ^{d'}
\\
\\
& {\scriptstyle{{\rm src}(F a)}}
\xlowertwocell[ddddl]{}^{<2>F v}{^{\psi}}
\xuppertwocell[ddddr]{}^{<2>F u}{^{\phi}}
\ar[dd] ^{e'}
\\
\\
& \scriptstyle{a'} \xtwocell[ddd]{}\omit{^{\omega'}}
\ar[ddl] _{v'}
\ar[ddr] ^{u'}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
\qquad = \qquad
\xy/33pt/
\xymatrix{
& \scriptstyle{a'} \xtwocell[ddd]{}\omit{^{\omega'}}
\ar[ddl] _{v'}
\ar[ddr] ^{u'}
\\
\\
\scriptstyle{{\rm trg}~(F r)} & & \scriptstyle{{\rm src}~)(F r)} \ar[ll] ^{F r}
\\
&
}
\endxy
$$
\<close>
proof -
have "(F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D
(\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>' =
(F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<Phi> (r, u) \<star>\<^sub>D d')) \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using D.whisker_right cmp_in_hom D.comp_assoc
by (metis D.arrI F\<omega> F\<omega>_def e'.ide_right)
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "(D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<Phi> (r, u) \<star>\<^sub>D d') =
D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u) \<star>\<^sub>D d'"
using cmp_in_hom cmp_components_are_iso D.whisker_right
by (metis C.hseqI D.comp_arr_inv' D.in_homE D.invert_opposite_sides_of_square
D.iso_inv_iso T.ide_base T.base_in_hom(1) \<open>trg\<^sub>C u = src\<^sub>C r\<close> e'.ide_right
preserves_arr u)
also have "... = (F r \<star>\<^sub>D F u) \<star>\<^sub>D d'"
using u cmp_components_are_iso D.comp_inv_arr' by simp
finally have "(F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, u)) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<Phi> (r, u) \<star>\<^sub>D d')) \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>' =
(F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F u) \<star>\<^sub>D d') \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
by simp
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F u) \<star>\<^sub>D d')) \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using D.comp_assoc by auto
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using u D.comp_arr_dom by simp
finally show ?thesis by blast
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<phi>) \<star>\<^sub>D d')) \<cdot>\<^sub>D (\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d') \<cdot>\<^sub>D
((\<omega>' \<star>\<^sub>D e') \<star>\<^sub>D d') \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have
"(F r \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e') \<cdot>\<^sub>D \<psi> \<star>\<^sub>D d' =
((F r \<star>\<^sub>D \<phi>) \<star>\<^sub>D d') \<cdot>\<^sub>D (\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d') \<cdot>\<^sub>D ((\<omega>' \<star>\<^sub>D e') \<star>\<^sub>D d') \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D d')"
using D.whisker_right \<phi> \<phi>' e' e'.antipar(1) u' u'_in_hhom
by (metis D.arrI D.seqE F\<omega> F\<omega>_def e'.ide_right)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d'] \<cdot>\<^sub>D
((\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d') \<cdot>\<^sub>D ((\<omega>' \<star>\<^sub>D e') \<star>\<^sub>D d')) \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "\<a>\<^sub>D[F r, F u, d'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<phi>) \<star>\<^sub>D d') =
(F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d']"
using D.assoc_naturality [of "F r" \<phi> d'] \<phi> by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d'] \<cdot>\<^sub>D
((\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d') \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D u', e', d'] \<cdot>\<^sub>D
(\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[v', e', d'])) \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
using F\<omega> F\<omega>_def \<omega>' D.comp_assoc D.hcomp_reassoc(1) [of \<omega>' e' d']
by (elim D.in_homE, simp)
also have "... = (F r \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D
(\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<psi>'"
proof -
have "D.seq (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d'])
(\<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d'] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d'))"
using u' by simp
moreover have "(F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d'] \<cdot>\<^sub>D
(\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d') =
\<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] \<cdot>\<^sub>D \<a>\<^sub>D[F r \<star>\<^sub>D u', e', d']"
using u' D.pentagon by simp
moreover have "D.iso (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d'])"
using u' by simp
moreover have "D.inv (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) = F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']"
using u' by simp
ultimately
have "\<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d'] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D u', e', d'] =
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d']"
using u' D.comp_assoc
D.invert_opposite_sides_of_square
[of "F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']"
"\<a>\<^sub>D[F r, u' \<star>\<^sub>D e', d'] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, u', e'] \<star>\<^sub>D d')"
"\<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d']" "\<a>\<^sub>D[F r \<star>\<^sub>D u', e', d']"]
by simp
thus ?thesis
using D.comp_assoc by metis
qed
also have
"... = (F r \<star>\<^sub>D \<r>\<^sub>D[u'] \<cdot>\<^sub>D (u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[u', e', d'] \<cdot>\<^sub>D (D.inv \<phi> \<star>\<^sub>D d')) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] \<cdot>\<^sub>D
(\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D (D.inv \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d'] \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
unfolding \<phi>'_def \<psi>'_def e'.trnr\<^sub>\<epsilon>_def d'.trnr\<^sub>\<eta>_def by simp
also have
"... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D
(F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D (D.inv \<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d'] \<cdot>\<^sub>D
(v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "F r \<star>\<^sub>D \<r>\<^sub>D[u'] \<cdot>\<^sub>D (u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[u', e', d'] \<cdot>\<^sub>D (D.inv \<phi> \<star>\<^sub>D d') =
(F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D
(F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d')"
proof -
have "D.ide (F r)" by simp
moreover have "D.seq \<r>\<^sub>D[u'] ((u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[u', e', d'] \<cdot>\<^sub>D (D.inv \<phi> \<star>\<^sub>D d')) \<and>
D.seq (u' \<star>\<^sub>D \<epsilon>') (\<a>\<^sub>D[u', e', d'] \<cdot>\<^sub>D (D.inv \<phi> \<star>\<^sub>D d')) \<and>
D.seq \<a>\<^sub>D[u', e', d'] (D.inv \<phi> \<star>\<^sub>D d')"
using \<phi>' \<phi>'_def unfolding e'.trnr\<^sub>\<epsilon>_def by blast
ultimately show ?thesis
using D.whisker_left by metis
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D
(((F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d')) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d'])) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D
\<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D (((\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D (D.inv \<psi> \<star>\<^sub>D d')) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d']) \<cdot>\<^sub>D
(v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
using D.comp_assoc by simp
also have
"... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D
((\<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d']) \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>')) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "((F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) =
F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']"
proof -
have "(F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') = F r \<star>\<^sub>D D.inv \<phi> \<cdot>\<^sub>D \<phi> \<star>\<^sub>D d'"
using u u' \<phi> 1 2 D.src_dom e'.antipar D.whisker_left D.whisker_right
by auto
also have "... = F r \<star>\<^sub>D (u' \<star>\<^sub>D e') \<star>\<^sub>D d'"
using \<phi> D.comp_inv_arr' by auto
finally have
"(F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d') = F r \<star>\<^sub>D (u' \<star>\<^sub>D e') \<star>\<^sub>D d'"
by simp
hence
"((F r \<star>\<^sub>D D.inv \<phi> \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<phi> \<star>\<^sub>D d')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']) =
(F r \<star>\<^sub>D (u' \<star>\<^sub>D e') \<star>\<^sub>D d') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d'])"
using D.comp_assoc by simp
also have "... = F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']"
proof -
have "\<guillemotleft>F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d'] :
F r \<star>\<^sub>D u' \<star>\<^sub>D e' \<star>\<^sub>D d' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D (u' \<star>\<^sub>D e') \<star>\<^sub>D d'\<guillemotright>"
using u' e'.antipar \<phi>' D.assoc'_in_hom
unfolding e'.trnr\<^sub>\<epsilon>_def
by (intro D.hcomp_in_vhom, auto)
thus ?thesis
using D.comp_cod_arr by blast
qed
finally show ?thesis by simp
qed
moreover have
"((\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D (D.inv \<psi> \<star>\<^sub>D d')) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d'] = \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d']"
proof -
have "(\<psi> \<star>\<^sub>D d') \<cdot>\<^sub>D (D.inv \<psi> \<star>\<^sub>D d') = (v' \<star>\<^sub>D e') \<star>\<^sub>D d'"
using \<psi> e'.antipar D.src_cod v' e'.antipar \<psi>' d'.trnr\<^sub>\<eta>_def
D.whisker_right [of d' \<psi> "D.inv \<psi>"] D.comp_arr_inv'
by auto
moreover have "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d'] : v' \<star>\<^sub>D e' \<star>\<^sub>D d' \<Rightarrow>\<^sub>D (v' \<star>\<^sub>D e') \<star>\<^sub>D d'\<guillemotright>"
using v' e'.antipar \<psi>' D.assoc'_in_hom
unfolding d'.trnr\<^sub>\<eta>_def
by fastforce
ultimately show ?thesis
using D.comp_cod_arr by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (((F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d'])) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d']) \<cdot>\<^sub>D
(\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "(\<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d']) \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') = v' \<star>\<^sub>D D.inv \<epsilon>'"
proof -
have 1: "D.hseq v' e'"
using v' e'.antipar \<psi>' unfolding d'.trnr\<^sub>\<eta>_def by fastforce
have "\<a>\<^sub>D[v', e', d'] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[v', e', d'] = v' \<star>\<^sub>D e' \<star>\<^sub>D d'"
using v' e'.antipar 1 D.comp_assoc_assoc' by auto
moreover have "\<guillemotleft>v' \<star>\<^sub>D D.inv \<epsilon>' : v' \<star>\<^sub>D trg\<^sub>D e' \<Rightarrow>\<^sub>D v' \<star>\<^sub>D e' \<star>\<^sub>D d'\<guillemotright>"
using v' e'.antipar 1
apply (intro D.hcomp_in_vhom)
apply auto
by (metis D.ideD(1) D.trg_src \<open>trg\<^sub>D e' = a'\<close> e'.antipar(2) e'.ide_right)
ultimately show ?thesis
using D.comp_cod_arr by auto
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D ((F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d']) \<cdot>\<^sub>D
(\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "((F r \<star>\<^sub>D \<a>\<^sub>D[u', e', d']) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d'])) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] =
\<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d']"
using \<phi> u' e'.antipar 1 D.comp_cod_arr D.comp_assoc_assoc'
D.whisker_left [of "F r" "\<a>\<^sub>D[u', e', d']" "\<a>\<^sub>D\<^sup>-\<^sup>1[u', e', d']"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', trg\<^sub>D e'] \<cdot>\<^sub>D (((F r \<star>\<^sub>D u') \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D
(\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d')) \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "(F r \<star>\<^sub>D u' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', e' \<star>\<^sub>D d'] =
\<a>\<^sub>D[F r, u', trg\<^sub>D e'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D u') \<star>\<^sub>D \<epsilon>')"
using D.assoc_naturality [of "F r" u' \<epsilon>'] e' u' u'_in_hhom by force
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', trg\<^sub>D e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D trg\<^sub>D e') \<cdot>\<^sub>D
((v' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>')) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "((F r \<star>\<^sub>D u') \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') = (\<omega>' \<star>\<^sub>D trg\<^sub>D e') \<cdot>\<^sub>D (v' \<star>\<^sub>D \<epsilon>')"
proof -
have "((F r \<star>\<^sub>D u') \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D e' \<star>\<^sub>D d') =
((F r \<star>\<^sub>D u') \<cdot>\<^sub>D \<omega>' \<star>\<^sub>D \<epsilon>' \<cdot>\<^sub>D (e' \<star>\<^sub>D d'))"
using D.interchange
by (metis D.comp_arr_dom D.hcomp_simps(3) D.hseqI D.ide_char D.in_hhomE
D.in_homE D.seqI T'.base_in_hom(1) T'.base_simps(3) T.base_simps(2)
\<omega>' e'.counit_simps(1) e'.counit_simps(2) preserves_src u' u'_in_hhom)
also have "... = \<omega>' \<cdot>\<^sub>D v' \<star>\<^sub>D trg\<^sub>D e' \<cdot>\<^sub>D \<epsilon>'"
using \<omega>' D.comp_arr_dom D.comp_cod_arr by auto
also have "... = (\<omega>' \<star>\<^sub>D trg\<^sub>D e') \<cdot>\<^sub>D (v' \<star>\<^sub>D \<epsilon>')"
using D.interchange
by (metis D.arrI D.comp_cod_arr D.ide_char D.seqI \<omega>' \<open>trg\<^sub>D e' = a'\<close>
e'.counit_simps(1) e'.counit_simps(3) e'.counit_simps(5) v' v'_def)
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u', trg\<^sub>D e'] \<cdot>\<^sub>D (\<omega>' \<star>\<^sub>D trg\<^sub>D e') \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[v']"
proof -
have "(v' \<star>\<^sub>D \<epsilon>') \<cdot>\<^sub>D (v' \<star>\<^sub>D D.inv \<epsilon>') = v' \<star>\<^sub>D trg\<^sub>D e'"
using v' D.whisker_left D.comp_arr_inv D.inv_is_inverse
by (metis D.comp_arr_inv' D.seqI' d'.unit_in_vhom e'.counit_in_hom(2)
e'.counit_is_iso e'.counit_simps(3))
moreover have "\<guillemotleft>\<r>\<^sub>D\<^sup>-\<^sup>1[v'] : v' \<Rightarrow>\<^sub>D v' \<star>\<^sub>D trg\<^sub>D e'\<guillemotright>"
using v' 1 by simp
ultimately show ?thesis
using v' D.comp_cod_arr by auto
qed
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (\<a>\<^sub>D[F r, u', trg\<^sub>D e'] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D u']) \<cdot>\<^sub>D \<omega>'"
using u' v' \<omega>' D.runit'_naturality D.comp_assoc
by (metis D.in_hhomE D.in_homE a'_def e')
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[u']) \<cdot>\<^sub>D \<omega>'"
using 1 T'.ide_base u' D.runit_hcomp [of "F r" u'] by fastforce
also have "... = ((F r \<star>\<^sub>D \<r>\<^sub>D[u']) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[u'])) \<cdot>\<^sub>D \<omega>'"
using D.comp_assoc by simp
also have "... = (F r \<star>\<^sub>D \<r>\<^sub>D[u'] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[u']) \<cdot>\<^sub>D \<omega>'"
using 1 T'.ide_base u' D.whisker_left by simp
also have "... = (F r \<star>\<^sub>D u') \<cdot>\<^sub>D \<omega>'"
using u'
by (metis D.comp_ide_self D.ide_in_hom(2) D.ide_is_iso
D.invert_opposite_sides_of_square D.invert_side_of_triangle(1)
D.iso_runit D.runit_in_vhom D.seqI')
also have "... = \<omega>'"
using \<omega>' D.comp_cod_arr by auto
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
thus "\<exists>w' \<theta>' \<nu>'. D.ide w' \<and> \<guillemotleft>\<theta>' : F f \<star>\<^sub>D w' \<Rightarrow>\<^sub>D u'\<guillemotright> \<and>
\<guillemotleft>\<nu>' : D.dom \<omega>' \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w'\<guillemotright> \<and> D.iso \<nu>' \<and> T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<nu>' = \<omega>'"
using w' \<theta>' \<nu>' iso_\<nu>' v'_def by blast
qed
text \<open>Now we establish \<open>T'.T2\<close>.\<close>
next
fix u w w' \<theta> \<theta>' \<beta>
assume w: "D.ide w"
assume w': "D.ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : F f \<star>\<^sub>D w \<Rightarrow>\<^sub>D u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : F f \<star>\<^sub>D w' \<Rightarrow>\<^sub>D u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : F g \<star>\<^sub>D w \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w'\<guillemotright>"
assume eq: "T'.composite_cell w \<theta> = T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<beta>"
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow>\<^sub>D w'\<guillemotright> \<and> \<beta> = F g \<star>\<^sub>D \<gamma> \<and> \<theta> = \<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>)"
proof -
define a where "a = src\<^sub>D w"
have a: "D.obj a"
unfolding a_def by (simp add: w)
have [simp]: "src\<^sub>D \<theta> = a"
using \<theta> a_def
by (metis D.dom_src D.in_homE D.src.preserves_dom D.src.preserves_reflects_arr
D.src_hcomp)
have [simp]: "trg\<^sub>D \<theta> = trg\<^sub>D (F f)"
using \<theta>
by (metis D.arr_dom D.in_homE D.trg_hcomp D.vconn_implies_hpar(2))
have [simp]: "src\<^sub>D \<theta>' = a"
using \<theta>' a_def
by (metis D.horizontal_homs_axioms D.in_homE \<open>src\<^sub>D \<theta> = a\<close> \<theta> horizontal_homs.src_cod)
have [simp]: "trg\<^sub>D \<theta>' = trg\<^sub>D (F f)"
using \<theta>'
by (metis D.vconn_implies_hpar(2) D.vconn_implies_hpar(4) \<open>trg\<^sub>D \<theta> = trg\<^sub>D (F f)\<close> \<theta>)
have [simp]: "src\<^sub>D w = a"
using a_def by simp
have [simp]: "trg\<^sub>D w = map\<^sub>0 (src\<^sub>C \<rho>)"
by (metis D.horizontal_homs_axioms D.hseq_char D.in_homE T.tab_simps(2) T.leg0_simps(2)
\<theta> category.ideD(1) category.ide_dom horizontal_homs_def preserves_src)
have [simp]: "src\<^sub>D w' = a"
using a_def
by (metis D.ideD(1) D.in_homE D.src_hcomp D.vconn_implies_hpar(1) \<open>src\<^sub>D \<theta>' = a\<close>
\<theta>' category.ide_dom horizontal_homs_def weak_arrow_of_homs_axioms
weak_arrow_of_homs_def)
have [simp]: "trg\<^sub>D w' = map\<^sub>0 (src\<^sub>C \<rho>)"
by (metis D.horizontal_homs_axioms D.hseq_char D.in_homE T.tab_simps(2) T.leg0_simps(2)
\<theta>' category.ideD(1) category.ide_dom horizontal_homs_def preserves_src)
text \<open>First, reflect the picture back to \<open>C\<close>, so that we will be able to apply \<open>T.T2\<close>.
We need to choose arrows in \<open>C\<close> carefully, so that their \<open>F\<close> images will enable the
cancellation of the various isomorphisms that appear.\<close>
obtain a\<^sub>C where a\<^sub>C: "C.obj a\<^sub>C \<and> D.equivalent_objects (map\<^sub>0 a\<^sub>C) a"
using w a_def biessentially_surjective_on_objects D.obj_src D.ideD(1)
by presburger
obtain e where e: "\<guillemotleft>e : map\<^sub>0 a\<^sub>C \<rightarrow>\<^sub>D a\<guillemotright> \<and> D.equivalence_map e"
using a\<^sub>C D.equivalent_objects_def by auto
obtain d \<eta> \<epsilon>
where d\<eta>\<epsilon>: "adjoint_equivalence_in_bicategory (\<cdot>\<^sub>D) (\<star>\<^sub>D) \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D e d \<eta> \<epsilon>"
using e D.equivalence_map_extends_to_adjoint_equivalence by blast
interpret e: adjoint_equivalence_in_bicategory \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D e d \<eta> \<epsilon>
using d\<eta>\<epsilon> by auto
interpret d: adjoint_equivalence_in_bicategory \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
d e "D.inv \<epsilon>" "D.inv \<eta>"
using e.dual_adjoint_equivalence by simp
have [simp]: "src\<^sub>D e = map\<^sub>0 a\<^sub>C"
using e by auto
have [simp]: "trg\<^sub>D e = a"
using e by auto
have [simp]: "src\<^sub>D d = a"
using e.antipar by simp
have [simp]: "trg\<^sub>D d = map\<^sub>0 a\<^sub>C"
using e.antipar by simp
have we: "\<guillemotleft>w \<star>\<^sub>D e : map\<^sub>0 a\<^sub>C \<rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C \<rho>)\<guillemotright>"
using a\<^sub>C e D.ideD(1) \<open>trg\<^sub>D w = map\<^sub>0 (src\<^sub>C \<rho>)\<close> a_def by blast
obtain w\<^sub>C where
w\<^sub>C: "C.ide w\<^sub>C \<and> \<guillemotleft>w\<^sub>C : a\<^sub>C \<rightarrow>\<^sub>C src\<^sub>C \<rho>\<guillemotright> \<and> D.isomorphic (F w\<^sub>C) (w \<star>\<^sub>D e)"
using a\<^sub>C e we locally_essentially_surjective [of a\<^sub>C "src\<^sub>C \<rho>" "w \<star>\<^sub>D e"]
C.obj_src T.tab_simps(1) e.ide_left w by blast
have w'e: "\<guillemotleft>w' \<star>\<^sub>D e : map\<^sub>0 a\<^sub>C \<rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C \<rho>)\<guillemotright>"
using a\<^sub>C e D.ideD(1) \<open>trg\<^sub>D w' = map\<^sub>0 (src\<^sub>C \<rho>)\<close> a_def \<open>src\<^sub>D w' = a\<close> w' by blast
obtain w\<^sub>C' where
w\<^sub>C': "C.ide w\<^sub>C' \<and> \<guillemotleft>w\<^sub>C' : a\<^sub>C \<rightarrow>\<^sub>C src\<^sub>C \<rho>\<guillemotright> \<and> D.isomorphic (F w\<^sub>C') (w' \<star>\<^sub>D e)"
using a\<^sub>C e a_def locally_essentially_surjective
by (metis C.obj_src D.ide_hcomp D.hseq_char D.in_hhomE T.tab_simps(2)
T.leg0_simps(2) e.ide_left w' w'e)
have [simp]: "src\<^sub>C w\<^sub>C = a\<^sub>C"
using w\<^sub>C by auto
have [simp]: "trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>"
using w\<^sub>C by auto
have [simp]: "src\<^sub>C w\<^sub>C' = a\<^sub>C"
using w\<^sub>C' by auto
have [simp]: "trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>"
using w\<^sub>C' by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : F w\<^sub>C \<Rightarrow>\<^sub>D w \<star>\<^sub>D e\<guillemotright> \<and> D.iso \<phi>"
using w\<^sub>C D.isomorphicE by blast
obtain \<phi>' where \<phi>': "\<guillemotleft>\<phi>' : F w\<^sub>C' \<Rightarrow>\<^sub>D w' \<star>\<^sub>D e\<guillemotright> \<and> D.iso \<phi>'"
using w\<^sub>C' D.isomorphicE by blast
have ue: "\<guillemotleft>u \<star>\<^sub>D e : map\<^sub>0 a\<^sub>C \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C f)\<guillemotright> \<and> D.ide (u \<star>\<^sub>D e)"
using a\<^sub>C e \<theta> e.ide_left
by (intro conjI, auto)
obtain u\<^sub>C where
u\<^sub>C: "C.ide u\<^sub>C \<and> \<guillemotleft>u\<^sub>C : a\<^sub>C \<rightarrow>\<^sub>C trg\<^sub>C f\<guillemotright> \<and> D.isomorphic (F u\<^sub>C) (u \<star>\<^sub>D e)"
using a\<^sub>C e ue locally_essentially_surjective [of a\<^sub>C "trg\<^sub>C f" "u \<star>\<^sub>D e"] by auto
have [simp]: "src\<^sub>C u\<^sub>C = a\<^sub>C"
using u\<^sub>C by auto
have [simp]: "trg\<^sub>C u\<^sub>C = trg\<^sub>C f"
using u\<^sub>C by auto
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : u \<star>\<^sub>D e \<Rightarrow>\<^sub>D F u\<^sub>C\<guillemotright> \<and> D.iso \<psi>"
using u\<^sub>C D.isomorphic_symmetric D.isomorphicE by blast
define F\<theta>\<^sub>C where
"F\<theta>\<^sub>C = \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
have 1: "\<guillemotleft>F\<theta>\<^sub>C : F (f \<star>\<^sub>C w\<^sub>C) \<Rightarrow>\<^sub>D F u\<^sub>C\<guillemotright>"
proof (unfold F\<theta>\<^sub>C_def, intro D.comp_in_homI)
show "\<guillemotleft>D.inv (\<Phi> (f, w\<^sub>C)) : F (f \<star>\<^sub>C w\<^sub>C) \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F w\<^sub>C\<guillemotright>"
by (simp add: cmp_in_hom(2) w\<^sub>C)
show "\<guillemotleft>F f \<star>\<^sub>D \<phi> : F f \<star>\<^sub>D F w\<^sub>C \<Rightarrow>\<^sub>D F f \<star>\<^sub>D w \<star>\<^sub>D e\<guillemotright>"
using w w\<^sub>C \<phi> by (intro D.hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] : F f \<star>\<^sub>D w \<star>\<^sub>D e \<Rightarrow>\<^sub>D (F f \<star>\<^sub>D w) \<star>\<^sub>D e\<guillemotright>"
using w D.assoc'_in_hom by simp
show "\<guillemotleft>\<theta> \<star>\<^sub>D e : (F f \<star>\<^sub>D w) \<star>\<^sub>D e \<Rightarrow>\<^sub>D u \<star>\<^sub>D e\<guillemotright>"
using w \<theta> by (intro D.hcomp_in_vhom, auto)
show "\<guillemotleft>\<psi> : u \<star>\<^sub>D e \<Rightarrow>\<^sub>D F u\<^sub>C\<guillemotright>"
using \<psi> by simp
qed
have 2: "\<exists>\<theta>\<^sub>C. \<guillemotleft>\<theta>\<^sub>C : f \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C u\<^sub>C\<guillemotright> \<and> F \<theta>\<^sub>C = F\<theta>\<^sub>C"
using u\<^sub>C w\<^sub>C 1 e \<theta> \<phi> locally_full by simp
obtain \<theta>\<^sub>C where \<theta>\<^sub>C: "\<guillemotleft>\<theta>\<^sub>C : f \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C u\<^sub>C\<guillemotright> \<and> F \<theta>\<^sub>C = F\<theta>\<^sub>C"
using 2 by auto
define F\<theta>\<^sub>C' where
"F\<theta>\<^sub>C' = \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C'))"
have 1: "\<guillemotleft>F\<theta>\<^sub>C' : F (f \<star>\<^sub>C w\<^sub>C') \<Rightarrow>\<^sub>D F u\<^sub>C\<guillemotright>"
proof (unfold F\<theta>\<^sub>C'_def, intro D.comp_in_homI)
show "\<guillemotleft>D.inv (\<Phi> (f, w\<^sub>C')) : F (f \<star>\<^sub>C w\<^sub>C') \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F w\<^sub>C'\<guillemotright>"
by (simp add: cmp_in_hom(2) w\<^sub>C')
show "\<guillemotleft>F f \<star>\<^sub>D \<phi>' : F f \<star>\<^sub>D F w\<^sub>C' \<Rightarrow>\<^sub>D F f \<star>\<^sub>D w' \<star>\<^sub>D e\<guillemotright>"
using w' w\<^sub>C' \<phi>' by (intro D.hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] : F f \<star>\<^sub>D w' \<star>\<^sub>D e \<Rightarrow>\<^sub>D (F f \<star>\<^sub>D w') \<star>\<^sub>D e\<guillemotright>"
using w' D.assoc'_in_hom by simp
show "\<guillemotleft>\<theta>' \<star>\<^sub>D e : (F f \<star>\<^sub>D w') \<star>\<^sub>D e \<Rightarrow>\<^sub>D u \<star>\<^sub>D e\<guillemotright>"
using w' \<theta>' by (intro D.hcomp_in_vhom, auto)
show "\<guillemotleft>\<psi> : u \<star>\<^sub>D e \<Rightarrow>\<^sub>D F u\<^sub>C\<guillemotright>"
using \<psi> by simp
qed
have 2: "\<exists>\<theta>\<^sub>C'. \<guillemotleft>\<theta>\<^sub>C' : f \<star>\<^sub>C w\<^sub>C' \<Rightarrow>\<^sub>C u\<^sub>C\<guillemotright> \<and> F \<theta>\<^sub>C' = F\<theta>\<^sub>C'"
using u\<^sub>C w\<^sub>C' 1 e \<theta> \<phi> locally_full by simp
obtain \<theta>\<^sub>C' where \<theta>\<^sub>C': "\<guillemotleft>\<theta>\<^sub>C' : f \<star>\<^sub>C w\<^sub>C' \<Rightarrow>\<^sub>C u\<^sub>C\<guillemotright> \<and> F \<theta>\<^sub>C' = F\<theta>\<^sub>C'"
using 2 by auto
define F\<beta>\<^sub>C where
"F\<beta>\<^sub>C = \<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
have F\<beta>\<^sub>C: "\<guillemotleft>F\<beta>\<^sub>C: F (g \<star>\<^sub>C w\<^sub>C) \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C w\<^sub>C')\<guillemotright>"
proof (unfold F\<beta>\<^sub>C_def, intro D.comp_in_homI)
show "\<guillemotleft>D.inv (\<Phi> (g, w\<^sub>C)) : F (g \<star>\<^sub>C w\<^sub>C) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F w\<^sub>C\<guillemotright>"
by (simp add: cmp_in_hom(2) w\<^sub>C)
show "\<guillemotleft>F g \<star>\<^sub>D \<phi> : F g \<star>\<^sub>D F w\<^sub>C \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w \<star>\<^sub>D e\<guillemotright>"
using w\<^sub>C \<phi> apply (intro D.hcomp_in_vhom) by auto
show "\<guillemotleft>\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] : F g \<star>\<^sub>D w \<star>\<^sub>D e \<Rightarrow>\<^sub>D (F g \<star>\<^sub>D w) \<star>\<^sub>D e\<guillemotright>"
using w D.assoc'_in_hom by simp
show "\<guillemotleft>\<beta> \<star>\<^sub>D e : (F g \<star>\<^sub>D w) \<star>\<^sub>D e \<Rightarrow>\<^sub>D (F g \<star>\<^sub>D w') \<star>\<^sub>D e\<guillemotright>"
using w \<beta> apply (intro D.hcomp_in_vhom) by auto
show "\<guillemotleft>\<a>\<^sub>D[F g, w', e] : (F g \<star>\<^sub>D w') \<star>\<^sub>D e \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w' \<star>\<^sub>D e\<guillemotright>"
using w' e.antipar D.assoc_in_hom by simp
show "\<guillemotleft>F g \<star>\<^sub>D D.inv \<phi>' : F g \<star>\<^sub>D w' \<star>\<^sub>D e \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F w\<^sub>C'\<guillemotright>"
using w' w\<^sub>C' \<phi>' by (intro D.hcomp_in_vhom, auto)
show "\<guillemotleft>\<Phi> (g, w\<^sub>C') : F g \<star>\<^sub>D F w\<^sub>C' \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C w\<^sub>C')\<guillemotright>"
using w\<^sub>C' cmp_in_hom by simp
qed
have 1: "\<exists>\<beta>\<^sub>C. \<guillemotleft>\<beta>\<^sub>C : g \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C g \<star>\<^sub>C w\<^sub>C'\<guillemotright> \<and> F \<beta>\<^sub>C = F\<beta>\<^sub>C"
using w\<^sub>C w\<^sub>C' F\<beta>\<^sub>C locally_full by simp
obtain \<beta>\<^sub>C where \<beta>\<^sub>C: "\<guillemotleft>\<beta>\<^sub>C : g \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C g \<star>\<^sub>C w\<^sub>C'\<guillemotright> \<and> F \<beta>\<^sub>C = F\<beta>\<^sub>C"
using 1 by auto
text \<open>
The following is the main calculation that needs to be done, to permit us
to apply \<open>T.T2\<close>.
Once again, it started out looking simple, but once all the necessary
isomorphisms are thrown in it looks much more complicated.
\<close>
have *: "T.composite_cell w\<^sub>C \<theta>\<^sub>C = T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C"
proof -
have par: "C.par (T.composite_cell w\<^sub>C \<theta>\<^sub>C) (T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C)"
proof -
have "\<guillemotleft>T.composite_cell w\<^sub>C \<theta>\<^sub>C : g \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<^sub>C\<guillemotright>"
using w\<^sub>C \<theta>\<^sub>C T.composite_cell_in_hom by simp
moreover have "\<guillemotleft>T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C : g \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<^sub>C\<guillemotright>"
proof (intro C.comp_in_homI)
show "\<guillemotleft>\<beta>\<^sub>C : g \<star>\<^sub>C w\<^sub>C \<Rightarrow>\<^sub>C g \<star>\<^sub>C w\<^sub>C'\<guillemotright>"
using \<beta>\<^sub>C by simp
show "\<guillemotleft>\<rho> \<star>\<^sub>C w\<^sub>C' : g \<star>\<^sub>C w\<^sub>C' \<Rightarrow>\<^sub>C (r \<star>\<^sub>C f) \<star>\<^sub>C w\<^sub>C'\<guillemotright>"
using w\<^sub>C' by (intro C.hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sub>C[r, f, w\<^sub>C'] : (r \<star>\<^sub>C f) \<star>\<^sub>C w\<^sub>C' \<Rightarrow>\<^sub>C r \<star>\<^sub>C f \<star>\<^sub>C w\<^sub>C'\<guillemotright>"
using w\<^sub>C' C.assoc_in_hom by simp
show "\<guillemotleft>r \<star>\<^sub>C \<theta>\<^sub>C' : r \<star>\<^sub>C f \<star>\<^sub>C w\<^sub>C' \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<^sub>C\<guillemotright>"
using w\<^sub>C' \<theta>\<^sub>C' by (intro C.hcomp_in_vhom, auto)
qed
ultimately show ?thesis
by (metis C.in_homE)
qed
moreover have "F (T.composite_cell w\<^sub>C \<theta>\<^sub>C) = F (T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C)"
proof -
have "F (T.composite_cell w\<^sub>C \<theta>\<^sub>C) = F (r \<star>\<^sub>C \<theta>\<^sub>C) \<cdot>\<^sub>D F \<a>\<^sub>C[r, f, w\<^sub>C] \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w\<^sub>C)"
using par by auto
also have "... = (\<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C))) \<cdot>\<^sub>D
(\<Phi> (r, f \<star>\<^sub>C w\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C))) \<cdot>\<^sub>D
(\<Phi> (r \<star>\<^sub>C f, w\<^sub>C) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C)))"
proof -
have "src\<^sub>C f = trg\<^sub>C w\<^sub>C \<and> C.hseq r \<theta>\<^sub>C \<and> C.hseq \<rho> w\<^sub>C"
using par by auto
thus ?thesis
using w\<^sub>C \<theta>\<^sub>C preserves_assoc preserves_hcomp
by (metis C.ideD(2) C.ideD(3) C.in_homE T.ide_base T.ide_leg0 T.leg0_simps(3)
T.tab_simps(4) T.tab_simps(5))
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C) \<cdot>\<^sub>D (((D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C))) \<cdot>\<^sub>D
(\<Phi> (r, f \<star>\<^sub>C w\<^sub>C))) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C))) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D ((D.inv (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C))) \<cdot>\<^sub>D
(\<Phi> (r \<star>\<^sub>C f, w\<^sub>C)) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C)) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using D.comp_assoc by simp
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D ((F r \<star>\<^sub>D F \<theta>\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C))) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C] \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C)) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have
"(D.inv (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (r \<star>\<^sub>C f, w\<^sub>C)) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C) = F \<rho> \<star>\<^sub>D F w\<^sub>C"
using w\<^sub>C \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> D.comp_inv_arr' cmp_in_hom cmp_components_are_iso
D.comp_cod_arr
by simp
moreover have
"((D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C))) \<cdot>\<^sub>D (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C))) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C)) =
F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C)"
using w\<^sub>C D.comp_cod_arr D.comp_inv_arr' cmp_simps(1,4) C.VV.cod_simp
by auto
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C] \<cdot>\<^sub>D
(?\<rho>' \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "(F r \<star>\<^sub>D F \<theta>\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C)) = F r \<star>\<^sub>D F \<theta>\<^sub>C \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C)"
using \<theta>\<^sub>C w\<^sub>C D.whisker_left cmp_in_hom
by (metis C.hseqE C.seqE D.seqI' T'.ide_base T.tab_simps(2) T.ide_leg0
par preserves_hom)
moreover have "(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C) = ?\<rho>' \<star>\<^sub>D F w\<^sub>C"
using D.whisker_right by (simp add: w\<^sub>C)
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D
D.inv (\<Phi> (f, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C] \<cdot>\<^sub>D
(?\<rho>' \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using \<theta>\<^sub>C F\<theta>\<^sub>C_def D.comp_assoc by simp
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D
((F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C]) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "F r \<star>\<^sub>D \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D
D.inv (\<Phi> (f, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C) =
F r \<star>\<^sub>D \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>)"
using cmp_in_hom cmp_components_are_iso D.comp_arr_dom
by (metis C.arrI D.cod_inv D.comp_inv_arr' D.seqE F\<theta>\<^sub>C_def T.tab_simps(2)
T.ide_leg0 \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> \<theta>\<^sub>C preserves_arr w\<^sub>C)
also have "... = (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D
(F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>)"
using D.whisker_left
by (metis (no_types, lifting) C.in_homE D.comp_assoc D.seqE F\<theta>\<^sub>C_def T'.ide_base
\<theta>\<^sub>C preserves_arr)
finally have "F r \<star>\<^sub>D \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D
D.inv (\<Phi> (f, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C) =
(F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D
(F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, w \<star>\<^sub>D e] \<cdot>\<^sub>D (((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C)) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "(F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C] =
\<a>\<^sub>D[F r, F f, w \<star>\<^sub>D e] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>)"
using w\<^sub>C \<phi> \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> D.assoc_naturality [of "F r" "F f" \<phi>]
by (metis (mono_tags, lifting) C.ideD(1) D.in_homE D.vconn_implies_hpar(2)
T'.base_simps(2-4) T'.leg0_simps(2-5) T.leg0_simps(2)
T.tab_simps(2) preserves_src preserves_trg)
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, w \<star>\<^sub>D e]) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D w \<star>\<^sub>D e) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C) = ?\<rho>' \<star>\<^sub>D \<phi> \<cdot>\<^sub>D F w\<^sub>C"
using \<phi> D.interchange
by (metis D.comp_arr_dom D.comp_cod_arr D.in_homE T'.tab_simps(1,5))
also have "... = ?\<rho>' \<star>\<^sub>D (w \<star>\<^sub>D e) \<cdot>\<^sub>D \<phi>"
using \<phi> w\<^sub>C D.comp_arr_dom D.comp_cod_arr by auto
also have "... = (?\<rho>' \<star>\<^sub>D w \<star>\<^sub>D e) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>)"
using \<phi> D.interchange
by (metis D.comp_arr_ide D.comp_cod_arr D.in_homE D.seqI' T'.ide_leg1
T'.leg1_in_hom(2) T'.tab_in_vhom')
finally have
"((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C) = (?\<rho>' \<star>\<^sub>D w \<star>\<^sub>D e) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f \<star>\<^sub>D w, e]) \<cdot>\<^sub>D
(\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D F f, w, e] \<cdot>\<^sub>D
(?\<rho>' \<star>\<^sub>D w \<star>\<^sub>D e)) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "D.inv (F r \<star>\<^sub>D \<a>\<^sub>D[F f, w, e]) = F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using w by simp
moreover have "D.seq (F r \<star>\<^sub>D \<a>\<^sub>D[F f, w, e])
(\<a>\<^sub>D[F r, F f \<star>\<^sub>D w, e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e))"
using w by simp
moreover have
"(F r \<star>\<^sub>D \<a>\<^sub>D[F f, w, e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f \<star>\<^sub>D w, e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e) =
\<a>\<^sub>D[F r, F f, w \<star>\<^sub>D e] \<cdot>\<^sub>D \<a>\<^sub>D[F r \<star>\<^sub>D F f, w, e]"
using w D.pentagon by simp
ultimately
have "(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w \<star>\<^sub>D e] =
\<a>\<^sub>D[F r, F f \<star>\<^sub>D w, e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D F f, w, e]"
using w D.comp_assoc
D.invert_opposite_sides_of_square
[of "F r \<star>\<^sub>D \<a>\<^sub>D[F f, w, e]" "\<a>\<^sub>D[F r, F f \<star>\<^sub>D w, e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e)"
"\<a>\<^sub>D[F r, F f, w \<star>\<^sub>D e]" "\<a>\<^sub>D[F r \<star>\<^sub>D F f, w, e]"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D (((F r \<star>\<^sub>D \<theta>) \<star>\<^sub>D e) \<cdot>\<^sub>D
(\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e) \<cdot>\<^sub>D ((?\<rho>' \<star>\<^sub>D w) \<star>\<^sub>D e)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have
"(F r \<star>\<^sub>D \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f \<star>\<^sub>D w, e] = \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<theta>) \<star>\<^sub>D e)"
using D.assoc_naturality [of "F r" \<theta> e] \<theta> by auto
moreover have "\<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D F f, w, e] \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D w \<star>\<^sub>D e) =
((?\<rho>' \<star>\<^sub>D w) \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]"
using w we e.ide_left D.assoc'_naturality [of ?\<rho>' w e] by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D
(T'.composite_cell w \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "((F r \<star>\<^sub>D \<theta>) \<star>\<^sub>D e) \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w] \<star>\<^sub>D e) \<cdot>\<^sub>D ((?\<rho>' \<star>\<^sub>D w) \<star>\<^sub>D e) =
T'.composite_cell w \<theta> \<star>\<^sub>D e"
proof -
have "\<guillemotleft>T'.composite_cell w \<theta> : F g \<star>\<^sub>D w \<Rightarrow>\<^sub>D F r \<star>\<^sub>D u\<guillemotright>"
using w we \<theta> \<open>src\<^sub>D \<theta> = a\<close> \<open>trg\<^sub>D e = a\<close> T'.composite_cell_in_hom
by (metis D.ideD(1) D.ide_in_hom(1) D.not_arr_null D.seq_if_composable
T'.leg1_simps(3) T.leg1_simps(2-3) T.tab_simps(2)
\<open>trg\<^sub>D w = map\<^sub>0 (src\<^sub>C \<rho>)\<close> a_def preserves_src ue)
thus ?thesis
using D.whisker_right D.arrI by auto
qed
thus ?thesis
using D.comp_assoc by simp
qed
finally have L: "F (T.composite_cell w\<^sub>C \<theta>\<^sub>C) =
\<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D
(T'.composite_cell w \<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
by simp
have "F (T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C) =
F ((r \<star>\<^sub>C \<theta>\<^sub>C') \<cdot>\<^sub>C \<a>\<^sub>C[r, f, w\<^sub>C'] \<cdot>\<^sub>C (\<rho> \<star>\<^sub>C w\<^sub>C') \<cdot>\<^sub>C \<beta>\<^sub>C)"
using C.comp_assoc by simp
also have "... = F(r \<star>\<^sub>C \<theta>\<^sub>C') \<cdot>\<^sub>D F \<a>\<^sub>C[r, f, w\<^sub>C'] \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w\<^sub>C') \<cdot>\<^sub>D F \<beta>\<^sub>C"
using C.comp_assoc par by fastforce
also have "... = (\<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C') \<cdot>\<^sub>D D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C'))) \<cdot>\<^sub>D
(\<Phi> (r, f \<star>\<^sub>C w\<^sub>C') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C'))) \<cdot>\<^sub>D
(\<Phi> (r \<star>\<^sub>C f, w\<^sub>C') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C'))) \<cdot>\<^sub>D
\<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "C.hseq r \<theta>\<^sub>C' \<and> C.hseq \<rho> w\<^sub>C'"
using par by blast
thus ?thesis
using w\<^sub>C' \<theta>\<^sub>C' \<beta>\<^sub>C F\<beta>\<^sub>C_def preserves_assoc [of r f w\<^sub>C'] preserves_hcomp
by force
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C') \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C'))) \<cdot>\<^sub>D
(\<Phi> (r, f \<star>\<^sub>C w\<^sub>C')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C'))) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D ((D.inv (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C')) \<cdot>\<^sub>D
\<Phi> (r \<star>\<^sub>C f, w\<^sub>C')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C')) \<cdot>\<^sub>D ((D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D
\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using D.comp_assoc by simp
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C'] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "(D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C'))) \<cdot>\<^sub>D (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')) =
F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')"
proof -
have "D.seq (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C')) (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')) \<and>
D.arr (D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C'))) \<and>
D.dom (D.inv (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C'))) =
D.cod (\<Phi> (r, f \<star>\<^sub>C w\<^sub>C') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')))"
by (metis D.seqE calculation par preserves_arr)
thus ?thesis
using C.ide_hcomp C.ideD(1) C.trg_hcomp D.invert_side_of_triangle(1)
T.ide_base T.ide_leg0 T.leg0_simps(3) T.tab_simps(2) cmp_components_are_iso
\<open>trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>\<close> w\<^sub>C'
by presburger
qed
moreover have
"(D.inv (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (r \<star>\<^sub>C f, w\<^sub>C')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w\<^sub>C') =
F \<rho> \<star>\<^sub>D F w\<^sub>C'"
proof -
have "D.seq (F \<rho> \<star>\<^sub>D F w\<^sub>C') (D.inv (\<Phi> (C.dom \<rho>, C.dom w\<^sub>C'))) \<and>
D.arr (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C')) \<and>
D.dom (\<Phi> (r \<star>\<^sub>C f, w\<^sub>C')) =
D.cod ((F \<rho> \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D D.inv (\<Phi> (C.dom \<rho>, C.dom w\<^sub>C')))"
by (metis C.hseqI' C.ide_char D.seqE T.tab_simps(1) T.tab_simps(5)
\<open>trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>\<close> preserves_arr preserves_hcomp w\<^sub>C')
thus ?thesis
by (metis (no_types) C.ide_hcomp C.ide_char C.hcomp_simps(1)
D.cod_comp D.comp_inv_arr' D.seqE T.ide_base T.ide_leg0 T.leg0_simps(3)
T.tab_simps(2) cmp_components_are_iso D.comp_cod_arr
\<open>trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>\<close> w\<^sub>C')
qed
moreover have "(D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') =
F g \<star>\<^sub>D D.inv \<phi>'"
proof -
have "(D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') =
(F g \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')"
using w\<^sub>C' \<beta>\<^sub>C F\<beta>\<^sub>C_def cmp_components_are_iso D.comp_inv_arr' by simp
also have "... = F g \<star>\<^sub>D D.inv \<phi>'"
using D.comp_cod_arr [of "F g \<star>\<^sub>D D.inv \<phi>'" "F g \<star>\<^sub>D F w\<^sub>C'"]
by (metis D.cod_inv D.comp_null(2) D.hseq_char' D.in_homE
D.is_weak_composition T'.leg1_simps(6) \<phi>'
weak_composition.hcomp_simps\<^sub>W\<^sub>C(3))
finally show ?thesis by blast
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>\<^sub>C') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C'] \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using w\<^sub>C' D.whisker_right cmp_in_hom cmp_components_are_iso by simp
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
D.inv (\<Phi> (f, w\<^sub>C'))) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C'] \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using \<theta>\<^sub>C' F\<theta>\<^sub>C'_def by simp
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D (F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
(((F r \<star>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C'))) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C'))) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C']) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "F r \<star>\<^sub>D \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
D.inv (\<Phi> (f, w\<^sub>C')) =
(F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D
(F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F r \<star>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C')))"
using D.whisker_left cmp_in_hom cmp_components_are_iso
by (metis C.arrI D.src.preserves_reflects_arr D.src_vcomp D.vseq_implies_hpar(1)
F\<theta>\<^sub>C'_def T'.ide_base \<theta>\<^sub>C' preserves_arr)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D ((F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C']) \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "((F r \<star>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C'))) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w\<^sub>C'))) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w\<^sub>C'] =
\<a>\<^sub>D[F r, F f, F w\<^sub>C']"
using cmp_in_hom cmp_components_are_iso D.comp_cod_arr
D.whisker_left [of "F r" "D.inv (\<Phi> (f, w\<^sub>C'))" "\<Phi> (f, w\<^sub>C')"]
by (simp add: D.comp_inv_arr' w\<^sub>C')
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e] \<cdot>\<^sub>D
(((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "(F r \<star>\<^sub>D F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w\<^sub>C'] =
\<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>')"
using w\<^sub>C' \<phi>' D.assoc_naturality [of "F r" "F f" \<phi>']
by (metis C.ideD(1) D.dom_trg D.in_homE D.trg.preserves_dom
T'.leg0_simps(2-5) T'.base_simps(2-4) T.tab_simps(2) T.leg0_simps(2)
\<open>trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>\<close> preserves_src preserves_trg)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e] \<cdot>\<^sub>D
(?\<rho>' \<star>\<^sub>D w' \<star>\<^sub>D e) \<cdot>\<^sub>D (((F g \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e]) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "((F r \<star>\<^sub>D F f) \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D F w\<^sub>C') = (?\<rho>' \<star>\<^sub>D w' \<star>\<^sub>D e) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>')"
using \<phi>' D.interchange D.comp_arr_dom D.comp_cod_arr
by (metis D.in_homE T'.tab_in_hom(2))
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D
((F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e]) \<cdot>\<^sub>D
(?\<rho>' \<star>\<^sub>D w' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "((F g \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] = \<a>\<^sub>D[F g, w', e]"
proof -
have "((F g \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] =
(F g \<star>\<^sub>D w' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e]"
by (metis D.arr_inv D.cod_inv D.comp_arr_inv' D.in_homE D.seqI
D.whisker_left T'.ide_leg1 \<phi>')
also have "... = \<a>\<^sub>D[F g, w', e]"
using w' D.comp_cod_arr by simp
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f \<star>\<^sub>D w', e]) \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D F f, w', e] \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D w' \<star>\<^sub>D e)) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D
(\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "D.inv (F r \<star>\<^sub>D \<a>\<^sub>D[F f, w', e]) = F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]"
using w' by simp
moreover have "D.seq (F r \<star>\<^sub>D \<a>\<^sub>D[F f, w', e])
(\<a>\<^sub>D[F r, F f \<star>\<^sub>D w', e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e))"
using w' by simp
moreover have "D.iso (F r \<star>\<^sub>D \<a>\<^sub>D[F f, w', e])"
using w' by simp
moreover have "D.iso \<a>\<^sub>D[F r \<star>\<^sub>D F f, w', e]"
using w' by simp
moreover have "(F r \<star>\<^sub>D \<a>\<^sub>D[F f, w', e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f \<star>\<^sub>D w', e] \<cdot>\<^sub>D
(\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e) =
\<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e] \<cdot>\<^sub>D \<a>\<^sub>D[F r \<star>\<^sub>D F f, w', e]"
using w' D.pentagon by simp
ultimately
have "(F r \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e] =
\<a>\<^sub>D[F r, F f \<star>\<^sub>D w', e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D F f, w', e]"
using w' D.comp_assoc
D.invert_opposite_sides_of_square
[of "F r \<star>\<^sub>D \<a>\<^sub>D[F f, w', e]" "\<a>\<^sub>D[F r, F f \<star>\<^sub>D w', e] \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e)"
"\<a>\<^sub>D[F r, F f, w' \<star>\<^sub>D e]" "\<a>\<^sub>D[F r \<star>\<^sub>D F f, w', e]"]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D (((F r \<star>\<^sub>D \<theta>') \<star>\<^sub>D e) \<cdot>\<^sub>D
(\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e) \<cdot>\<^sub>D ((?\<rho>' \<star>\<^sub>D w') \<star>\<^sub>D e)) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w', e] \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e]) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "(F r \<star>\<^sub>D \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f \<star>\<^sub>D w', e] =
\<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<theta>') \<star>\<^sub>D e)"
using D.assoc_naturality [of "F r" \<theta>' e] \<theta>' by auto
moreover have "\<a>\<^sub>D\<^sup>-\<^sup>1[F r \<star>\<^sub>D F f, w', e] \<cdot>\<^sub>D (?\<rho>' \<star>\<^sub>D w' \<star>\<^sub>D e) =
((?\<rho>' \<star>\<^sub>D w') \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w', e]"
using w' w'e D.assoc'_naturality [of ?\<rho>' w' e] by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D
(T'.composite_cell w' \<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "((F r \<star>\<^sub>D \<theta>') \<star>\<^sub>D e) \<cdot>\<^sub>D (\<a>\<^sub>D[F r, F f, w'] \<star>\<^sub>D e) \<cdot>\<^sub>D ((?\<rho>' \<star>\<^sub>D w') \<star>\<^sub>D e) =
T'.composite_cell w' \<theta>' \<star>\<^sub>D e"
proof -
have "\<guillemotleft>T'.composite_cell w' \<theta>' : F g \<star>\<^sub>D w' \<Rightarrow>\<^sub>D F r \<star>\<^sub>D u\<guillemotright>"
using \<theta>' w' T'.composite_cell_in_hom D.vconn_implies_hpar(3) by simp
thus ?thesis
using D.whisker_right D.arrI by auto
qed
moreover have "(\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w', e] \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e]) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) = \<beta> \<star>\<^sub>D e"
using w' \<beta> e.ide_left \<open>src\<^sub>D w' = a\<close> \<open>trg\<^sub>D e = a\<close> F\<beta>\<^sub>C F\<beta>\<^sub>C_def D.comp_cod_arr
D.comp_arr_inv'
by (metis (no_types, lifting) D.comp_assoc_assoc'(2) D.hcomp_simps(1)
D.hcomp_simps(4) D.hseqI' D.ide_char D.in_homE D.vconn_implies_hpar(1)
D.vconn_implies_hpar(3) T'.ide_leg1 T.leg1_simps(2) T.leg1_simps(3)
T.tab_simps(2) \<open>trg\<^sub>D w' = map\<^sub>0 (src\<^sub>C \<rho>)\<close> preserves_src)
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D
(T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
proof -
have "D.arr (T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<beta>)"
by (metis (full_types) D.hseq_char D.seqE L eq par preserves_arr)
thus ?thesis
using D.whisker_right by (metis D.comp_assoc e.ide_left)
qed
finally have R: "F (T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C) =
\<Phi> (r, u\<^sub>C) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<psi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, u, e] \<cdot>\<^sub>D
(T'.composite_cell w' \<theta>' \<cdot>\<^sub>D \<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
by simp
show "F (T.composite_cell w\<^sub>C \<theta>\<^sub>C) = F (T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C)"
using eq L R by simp
qed
ultimately show ?thesis
using is_faithful [of "T.composite_cell w\<^sub>C \<theta>\<^sub>C" "T.composite_cell w\<^sub>C' \<theta>\<^sub>C' \<cdot>\<^sub>C \<beta>\<^sub>C"]
by simp
qed
have **: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>C \<Rightarrow>\<^sub>C w\<^sub>C'\<guillemotright> \<and> \<beta>\<^sub>C = g \<star>\<^sub>C \<gamma> \<and> \<theta>\<^sub>C = \<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)"
using * w\<^sub>C w\<^sub>C' \<theta>\<^sub>C \<theta>\<^sub>C' \<beta>\<^sub>C T.T2 [of w\<^sub>C w\<^sub>C' \<theta>\<^sub>C u\<^sub>C \<theta>\<^sub>C' \<beta>\<^sub>C] by simp
obtain \<gamma>\<^sub>C where
\<gamma>\<^sub>C: "\<guillemotleft>\<gamma>\<^sub>C : w\<^sub>C \<Rightarrow>\<^sub>C w\<^sub>C'\<guillemotright> \<and> \<beta>\<^sub>C = g \<star>\<^sub>C \<gamma>\<^sub>C \<and> \<theta>\<^sub>C = \<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C)"
using ** by auto
have \<gamma>\<^sub>C_unique: "\<And>\<gamma>\<^sub>C'. \<guillemotleft>\<gamma>\<^sub>C' : w\<^sub>C \<Rightarrow>\<^sub>C w\<^sub>C'\<guillemotright> \<and> \<beta>\<^sub>C = g \<star>\<^sub>C \<gamma>\<^sub>C' \<and>
\<theta>\<^sub>C = \<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C') \<Longrightarrow> \<gamma>\<^sub>C' = \<gamma>\<^sub>C"
using \<gamma>\<^sub>C ** by blast
text \<open>
Now use \<open>F\<close> to map everything back to \<open>D\<close>, transport the result along the
equivalence map \<open>e\<close>, and cancel all of the isomorphisms that got introduced.
\<close>
let ?P = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : w \<star>\<^sub>D e \<Rightarrow>\<^sub>D w' \<star>\<^sub>D e\<guillemotright> \<and>
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] = F g \<star>\<^sub>D \<gamma> \<and>
\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>)"
define \<gamma>e where "\<gamma>e = \<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi>"
have P\<gamma>e: "?P \<gamma>e"
proof -
have 1: "\<guillemotleft>F \<gamma>\<^sub>C : F w\<^sub>C \<Rightarrow>\<^sub>D F w\<^sub>C'\<guillemotright> \<and>
F \<beta>\<^sub>C = \<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C)) \<and>
F \<theta>\<^sub>C = F \<theta>\<^sub>C' \<cdot>\<^sub>D \<Phi> (f, C.cod \<gamma>\<^sub>C) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
using \<beta>\<^sub>C \<theta>\<^sub>C \<gamma>\<^sub>C preserves_hcomp [of f \<gamma>\<^sub>C] preserves_hcomp [of g \<gamma>\<^sub>C] by force
have A: "\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] =
F g \<star>\<^sub>D \<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi>"
proof -
have "F g \<star>\<^sub>D F \<gamma>\<^sub>C = D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D F \<beta>\<^sub>C \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C)"
proof -
have "F \<beta>\<^sub>C = \<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using 1 by simp
hence "D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D F \<beta>\<^sub>C = (F g \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using w\<^sub>C w\<^sub>C' \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> \<open>trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>\<close> cmp_components_are_iso
D.invert_side_of_triangle(1)
by (metis D.arrI F\<beta>\<^sub>C T.ide_leg1 T.leg1_simps(3) T.tab_simps(2) \<beta>\<^sub>C)
hence "(D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D F \<beta>\<^sub>C) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C) = F g \<star>\<^sub>D F \<gamma>\<^sub>C"
using cmp_components_are_iso D.invert_side_of_triangle(2)
by (metis "1" D.arrI D.inv_inv D.iso_inv_iso D.seqE F\<beta>\<^sub>C T.ide_leg1
T.leg1_simps(3) T.tab_simps(2) \<beta>\<^sub>C \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> w\<^sub>C)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C)"
using \<beta>\<^sub>C F\<beta>\<^sub>C_def D.comp_assoc by simp
also have "... = (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>)"
proof -
have "(D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') = F g \<star>\<^sub>D D.inv \<phi>'"
proof -
have "(D.inv (\<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') =
(F g \<star>\<^sub>D F w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>')"
using w\<^sub>C' \<phi>' cmp_components_are_iso D.comp_inv_arr' by simp
also have "... = F g \<star>\<^sub>D D.inv \<phi>'"
using w\<^sub>C' \<phi>' D.comp_cod_arr
by (metis D.arr_inv D.cod_inv D.in_homE D.whisker_left T'.ide_leg1)
finally show ?thesis by blast
qed
moreover have "(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C) = F g \<star>\<^sub>D \<phi>"
proof -
have "(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C)) \<cdot>\<^sub>D \<Phi> (g, w\<^sub>C) =
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D (F g \<star>\<^sub>D F w\<^sub>C)"
using w\<^sub>C \<phi> \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> cmp_components_are_iso cmp_in_hom
D.comp_inv_arr'
by simp
also have "... = F g \<star>\<^sub>D \<phi>"
using w\<^sub>C \<phi> D.comp_arr_dom
by (metis D.hcomp_simps(3) D.hseqI' D.in_hhom_def D.in_homE
D.vconn_implies_hpar(2) D.vconn_implies_hpar(4) T'.leg1_simps(2,5)
T.leg1_simps(2-3) T.tab_simps(2) preserves_src we)
finally show ?thesis by blast
qed
ultimately show ?thesis by simp
qed
finally have 2: "(F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D (\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) =
F g \<star>\<^sub>D F \<gamma>\<^sub>C"
using D.comp_assoc by simp
have 3: "(\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) =
(F g \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>C)"
proof -
have "D.hseq (F g) (F \<gamma>\<^sub>C)"
using "1" F\<beta>\<^sub>C \<beta>\<^sub>C by auto
moreover have "D.iso (F g \<star>\<^sub>D D.inv \<phi>')"
by (metis "2" D.iso_hcomp D.hseqE D.ide_is_iso D.iso_inv_iso D.seqE
T'.ide_leg1 \<phi>' calculation)
moreover have "D.inv (F g \<star>\<^sub>D D.inv \<phi>') = F g \<star>\<^sub>D \<phi>'"
by (metis D.hseqE D.ide_is_iso D.inv_hcomp D.inv_ide D.inv_inv D.iso_inv_iso
D.iso_is_arr T'.ide_leg1 \<phi>' calculation(2))
ultimately show ?thesis
using 2 \<phi> \<phi>'
D.invert_side_of_triangle(1)
[of "F g \<star>\<^sub>D F \<gamma>\<^sub>C" "F g \<star>\<^sub>D D.inv \<phi>'"
"(\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>)"]
by auto
qed
hence "\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] =
((F g \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>C)) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>)"
proof -
have "D.seq (F g \<star>\<^sub>D \<phi>') (F g \<star>\<^sub>D F \<gamma>\<^sub>C)"
by (metis "1" "2" "3" D.arrI D.comp_null(1) D.comp_null(2) D.ext F\<beta>\<^sub>C \<beta>\<^sub>C)
moreover have "D.iso (F g \<star>\<^sub>D \<phi>)"
using D.vconn_implies_hpar(2) D.vconn_implies_hpar(4) \<phi> we by auto
moreover have "D.inv (F g \<star>\<^sub>D \<phi>) = F g \<star>\<^sub>D D.inv \<phi>"
by (metis D.hseqE D.ide_is_iso D.inv_hcomp D.inv_ide D.iso_is_arr
T'.ide_leg1 \<phi> calculation(2))
ultimately show ?thesis
using 3 \<phi> \<phi>'
D.invert_side_of_triangle(2)
[of "(F g \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>C)"
"\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]" "F g \<star>\<^sub>D \<phi>"]
by auto
qed
also have "... = F g \<star>\<^sub>D \<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi>"
using \<phi>' D.whisker_left
by (metis "1" D.arr_inv D.cod_comp D.cod_inv D.comp_assoc D.in_homE D.seqI
T'.ide_leg1 \<phi>)
finally show ?thesis by simp
qed
have B: "\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi>)"
proof -
have "F \<theta>\<^sub>C' \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C)) =
(\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D
\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C)) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
using \<gamma>\<^sub>C \<theta>\<^sub>C' F\<theta>\<^sub>C'_def D.comp_assoc by auto
also have "... = \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
proof -
have "(D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) = F f \<star>\<^sub>D F \<gamma>\<^sub>C"
using D.comp_cod_arr
by (metis (mono_tags, lifting) C.in_homE D.cod_comp D.comp_inv_arr' D.seqE
T.tab_simps(2) T.ide_leg0 cmp_components_are_iso \<gamma>\<^sub>C 1 \<open>trg\<^sub>C w\<^sub>C' = src\<^sub>C \<rho>\<close>
\<theta>\<^sub>C preserves_arr w\<^sub>C')
thus ?thesis
using D.comp_assoc by simp
qed
finally have "F \<theta>\<^sub>C' \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C)) =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
by simp
hence 3: "F \<theta>\<^sub>C' \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C)"
using cmp_components_are_iso D.iso_inv_iso D.iso_is_retraction D.retraction_is_epi
D.epiE
by (metis C.in_homE D.comp_assoc T.tab_simps(2) T.ide_leg0 \<gamma>\<^sub>C 1
\<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> \<theta>\<^sub>C preserves_arr w\<^sub>C)
hence "(\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>)) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C)) =
(\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<gamma>\<^sub>C)) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
using 1 \<theta>\<^sub>C F\<theta>\<^sub>C_def D.comp_assoc by (metis C.in_homE \<gamma>\<^sub>C)
hence 2: "(\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C)"
using \<gamma>\<^sub>C cmp_components_are_iso D.iso_inv_iso D.iso_is_retraction
D.retraction_is_epi D.epiE
by (metis (mono_tags, lifting) 1 3 C.in_homE D.comp_assoc T.tab_simps(2)
T.ide_leg0 \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close> \<theta>\<^sub>C preserves_arr w\<^sub>C)
hence "\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] =
(\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<phi>)"
proof -
have "D.inv (F f \<star>\<^sub>D \<phi>) = F f \<star>\<^sub>D D.inv \<phi>"
using \<phi>
by (metis C.arrI D.hseq_char D.ide_is_iso D.inv_hcomp D.inv_ide D.seqE F\<theta>\<^sub>C_def
T'.ide_leg0 preserves_arr \<theta>\<^sub>C)
thus ?thesis
using \<phi> \<phi>' \<theta> \<theta>' \<gamma>\<^sub>C D.invert_side_of_triangle(2)
by (metis 2 C.arrI D.comp_assoc D.iso_hcomp D.hseqE D.ide_is_iso D.seqE
F\<theta>\<^sub>C_def T'.ide_leg0 \<theta>\<^sub>C preserves_arr)
qed
also have "... = \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>C) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<phi>)"
using D.comp_assoc by simp
also have
"... = \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi>)"
proof -
have "D.arr (\<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi>)"
using "1" \<phi> \<phi>' by blast
thus ?thesis
using D.whisker_left by auto
qed
finally show ?thesis by simp
qed
have C: "\<guillemotleft>\<phi>' \<cdot>\<^sub>D F \<gamma>\<^sub>C \<cdot>\<^sub>D D.inv \<phi> : w \<star>\<^sub>D e \<Rightarrow>\<^sub>D w' \<star>\<^sub>D e\<guillemotright>"
using \<phi> \<phi>' \<gamma>\<^sub>C 1 by (meson D.comp_in_homI D.inv_in_hom)
show ?thesis
unfolding \<gamma>e_def
using A B C by simp
qed
have UN: "\<And>\<gamma>. ?P \<gamma> \<Longrightarrow> \<gamma> = \<gamma>e"
proof -
fix \<gamma>
assume \<gamma>: "?P \<gamma>"
show "\<gamma> = \<gamma>e"
proof -
let ?\<gamma>' = "D.inv \<phi>' \<cdot>\<^sub>D \<gamma> \<cdot>\<^sub>D \<phi>"
have \<gamma>': "\<guillemotleft>?\<gamma>' : F w\<^sub>C \<Rightarrow>\<^sub>D F w\<^sub>C'\<guillemotright>"
using \<gamma> \<phi> \<phi>' by auto
obtain \<gamma>\<^sub>C' where \<gamma>\<^sub>C': "\<guillemotleft>\<gamma>\<^sub>C' : w\<^sub>C \<Rightarrow>\<^sub>C w\<^sub>C'\<guillemotright> \<and> F \<gamma>\<^sub>C' = ?\<gamma>'"
using w\<^sub>C w\<^sub>C' \<gamma> \<gamma>' locally_full by fastforce
have 1: "\<beta>\<^sub>C = g \<star>\<^sub>C \<gamma>\<^sub>C'"
proof -
have "F \<beta>\<^sub>C = F (g \<star>\<^sub>C \<gamma>\<^sub>C')"
proof -
have "F \<beta>\<^sub>C =
\<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using \<beta>\<^sub>C F\<beta>\<^sub>C_def by simp
have "F (g \<star>\<^sub>C \<gamma>\<^sub>C') =
\<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>' \<cdot>\<^sub>D \<gamma> \<cdot>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using \<gamma>\<^sub>C' preserves_hcomp
by (metis C.hseqI' C.in_homE C.trg_dom T.tab_simps(2) T.leg1_simps(2)
T.leg1_simps(3,5-6) \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close>)
also have "... = \<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using \<phi> \<phi>' D.whisker_left D.comp_assoc
by (metis D.arrI D.seqE F\<beta>\<^sub>C_def T'.ide_leg1 \<gamma> \<gamma>')
also have "... = \<Phi> (g, w\<^sub>C') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D
(\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (g, w\<^sub>C))"
using \<gamma> D.comp_assoc by simp
also have "... = F \<beta>\<^sub>C"
using \<beta>\<^sub>C F\<beta>\<^sub>C_def D.comp_assoc by simp
finally show ?thesis by simp
qed
moreover have "C.par \<beta>\<^sub>C (g \<star>\<^sub>C \<gamma>\<^sub>C')"
proof (intro conjI)
show "C.arr \<beta>\<^sub>C"
using \<beta>\<^sub>C by blast
show 2: "C.hseq g \<gamma>\<^sub>C'"
using F\<beta>\<^sub>C \<beta>\<^sub>C calculation by fastforce
show "C.dom \<beta>\<^sub>C = C.dom (g \<star>\<^sub>C \<gamma>\<^sub>C')"
using 2 \<beta>\<^sub>C \<gamma>\<^sub>C' by fastforce
show "C.cod \<beta>\<^sub>C = C.cod (g \<star>\<^sub>C \<gamma>\<^sub>C')"
using 2 \<beta>\<^sub>C \<gamma>\<^sub>C' by fastforce
qed
ultimately show ?thesis using is_faithful by blast
qed
have 2: "\<theta>\<^sub>C = \<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C')"
proof -
have "F \<theta>\<^sub>C = F (\<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C'))"
proof -
have "F (\<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C')) = F \<theta>\<^sub>C' \<cdot>\<^sub>D F (f \<star>\<^sub>C \<gamma>\<^sub>C')"
using \<theta>\<^sub>C' \<gamma>\<^sub>C' by force
also have
"... = F \<theta>\<^sub>C' \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<phi>' \<cdot>\<^sub>D \<gamma> \<cdot>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
using w\<^sub>C w\<^sub>C' \<theta>\<^sub>C' \<gamma>\<^sub>C' preserves_hcomp
by (metis C.hseqI' C.in_homE C.trg_dom T.tab_simps(2) T.leg0_simps(2)
T.leg0_simps(4-5) \<open>trg\<^sub>C w\<^sub>C = src\<^sub>C \<rho>\<close>)
also have "... = F \<theta>\<^sub>C' \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C') \<cdot>\<^sub>D
((F f \<star>\<^sub>D D.inv \<phi>') \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, w\<^sub>C))"
using D.whisker_left
by (metis D.arrI D.seqE T'.ide_leg0 \<gamma>')
also have "... = \<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (((F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D
(D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<gamma>)) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
using \<theta>\<^sub>C' F\<theta>\<^sub>C'_def D.comp_assoc by simp
also have "... = (\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>)) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
proof -
have "D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C') = F f \<star>\<^sub>D F w\<^sub>C'"
using w\<^sub>C' cmp_in_hom cmp_components_are_iso
by (simp add: D.comp_inv_arr')
moreover have "D.hseq (F f) (D.inv \<phi>')"
using \<phi>' D.hseqI'
by (metis D.ide_is_iso D.in_hhom_def D.iso_inv_iso D.iso_is_arr
D.trg_inv D.vconn_implies_hpar(2) D.vconn_implies_hpar(4)
T'.ide_leg0 T'.leg1_simps(3) T.leg1_simps(2-3)
T.tab_simps(2) \<gamma> preserves_src we)
ultimately have "(D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<phi>') =
F f \<star>\<^sub>D D.inv \<phi>'"
using w\<^sub>C' \<phi>' D.comp_cod_arr [of "F f \<star>\<^sub>D D.inv \<phi>'" "F f \<star>\<^sub>D F w\<^sub>C'"]
by fastforce
hence "((F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>) =
((F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>)"
by simp
also have "... = F f \<star>\<^sub>D \<gamma>"
using \<gamma> \<phi>' \<theta>\<^sub>C' F\<theta>\<^sub>C'_def D.comp_cod_arr D.whisker_left D.hseqI'
by (metis D.comp_arr_inv' D.in_hhom_def D.in_homE T'.ide_leg0 w'e)
finally have "((F f \<star>\<^sub>D \<phi>') \<cdot>\<^sub>D (D.inv (\<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D \<Phi> (f, w\<^sub>C')) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv \<phi>')) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>) =
F f \<star>\<^sub>D \<gamma>"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w\<^sub>C))"
using \<gamma> D.comp_assoc by metis
also have "... = F \<theta>\<^sub>C"
using \<theta>\<^sub>C F\<theta>\<^sub>C_def by simp
finally show ?thesis by simp
qed
moreover have "C.par \<theta>\<^sub>C (\<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C'))"
proof (intro conjI)
show "C.arr \<theta>\<^sub>C"
using \<theta>\<^sub>C by auto
show 1: "C.seq \<theta>\<^sub>C' (f \<star>\<^sub>C \<gamma>\<^sub>C')"
using \<theta>\<^sub>C' \<gamma>\<^sub>C'
by (metis C.arrI \<theta>\<^sub>C calculation preserves_reflects_arr)
show "C.dom \<theta>\<^sub>C = C.dom (\<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C'))"
using 1 \<theta>\<^sub>C \<gamma>\<^sub>C' by fastforce
show "C.cod \<theta>\<^sub>C = C.cod (\<theta>\<^sub>C' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>C'))"
using 1 \<theta>\<^sub>C \<gamma>\<^sub>C' \<gamma>\<^sub>C by auto
qed
ultimately show ?thesis
using is_faithful by blast
qed
have "F \<gamma>\<^sub>C' = F \<gamma>\<^sub>C"
using ** \<gamma>\<^sub>C \<gamma>\<^sub>C' 1 2 by blast
hence "?\<gamma>' = F \<gamma>\<^sub>C"
using \<gamma>\<^sub>C' by simp
thus "\<gamma> = \<gamma>e"
unfolding \<gamma>e_def
by (metis D.arrI D.comp_assoc D.inv_inv D.invert_side_of_triangle(1)
D.invert_side_of_triangle(2) D.iso_inv_iso \<gamma>' \<phi> \<phi>')
qed
qed
text \<open>We are now in a position to exhibit the 2-cell \<open>\<gamma>\<close> and show that it
is unique with the required properties.\<close>
show ?thesis
proof
let ?\<gamma> = "\<r>\<^sub>D[w'] \<cdot>\<^sub>D (w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[w', e, d] \<cdot>\<^sub>D (\<gamma>e \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D
(w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
have \<gamma>: "\<guillemotleft>?\<gamma> : w \<Rightarrow>\<^sub>D w'\<guillemotright>"
using P\<gamma>e w w' e.counit_in_hom(2) e.counit_is_iso
apply (intro D.comp_in_homI)
apply auto[2]
apply fastforce
apply auto[3]
apply fastforce
by auto
moreover have "\<beta> = F g \<star>\<^sub>D ?\<gamma>"
proof -
have "F g \<star>\<^sub>D ?\<gamma> =
(F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D (F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D (F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using w w' \<gamma> P\<gamma>e D.whisker_left e.antipar
by (metis D.arrI D.seqE T'.ide_leg1)
also have "... =
(F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D (F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F g \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D (F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F g \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d] =
\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using w w' e.antipar P\<gamma>e D.assoc'_naturality [of "F g" \<gamma>e d]
by (metis D.dom_trg D.ideD(1-3) D.in_hhomE D.in_homE
D.src_dom D.trg.preserves_dom T'.leg1_simps(2) T'.leg1_simps(3,5-6)
T.tab_simps(2) T.leg0_simps(2) e e.ide_right preserves_src we)
also have
"... = (\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using D.comp_assoc by simp
also have "... = F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
proof -
have "(\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) =
(F g \<star>\<^sub>D (w' \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using w'e D.isomorphic_implies_ide(2) w\<^sub>C' D.comp_assoc_assoc'(1) by auto
also have "... = F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
proof -
have "\<guillemotleft>F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d : F g \<star>\<^sub>D (w \<star>\<^sub>D e) \<star>\<^sub>D d \<Rightarrow>\<^sub>D F g \<star>\<^sub>D (w' \<star>\<^sub>D e) \<star>\<^sub>D d\<guillemotright>"
using we e.ide_right e.antipar P\<gamma>e by fastforce
thus ?thesis
using D.comp_cod_arr by auto
qed
finally show ?thesis by blast
qed
finally have
"\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F g \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d] =
F g \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
by simp
thus ?thesis by simp
qed
also have "... =
(F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D (F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D
(\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D (F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using P\<gamma>e by simp
also have
"... =
(F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D (F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (\<a>\<^sub>D[F g, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D
((\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d =
(\<a>\<^sub>D[F g, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D ((\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d)"
proof -
have "D.arr (\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e])"
using D.arrI D.in_hhom_def D.vconn_implies_hpar(2) P\<gamma>e we by auto
thus ?thesis
using D.whisker_right by auto
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... =
(F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D (F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D
((F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (\<a>\<^sub>D[F g, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w', e, d]) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D[F g \<star>\<^sub>D w, e, d]) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d])) \<cdot>\<^sub>D
(F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d =
\<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w', e, d] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D[F g \<star>\<^sub>D w, e, d]"
proof -
have "src\<^sub>D \<beta> = trg\<^sub>D e"
using \<beta>
by (metis D.dom_trg D.hseq_char' D.in_homE D.src_dom D.src_hcomp
D.trg.is_extensional D.trg.preserves_arr D.trg.preserves_dom
\<open>trg\<^sub>D e = a\<close> a_def)
moreover have "src\<^sub>D (F g) = trg\<^sub>D w"
by simp
moreover have "src\<^sub>D (F g) = trg\<^sub>D w'"
by simp
moreover have
"\<guillemotleft>(\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d : ((F g \<star>\<^sub>D w) \<star>\<^sub>D e) \<star>\<^sub>D d \<Rightarrow>\<^sub>D ((F g \<star>\<^sub>D w') \<star>\<^sub>D e) \<star>\<^sub>D d\<guillemotright>"
using \<beta> w w' e e.antipar
by (intro D.hcomp_in_vhom, auto)
ultimately have
"\<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w', e, d] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D[F g \<star>\<^sub>D w, e, d] =
\<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w', e, d] \<cdot>\<^sub>D \<a>\<^sub>D[F g \<star>\<^sub>D w', e, d] \<cdot>\<^sub>D ((\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d)"
using w' e e.ide_left e.ide_right e.antipar \<beta> D.assoc'_naturality
by (metis D.assoc_naturality D.in_homE e.triangle_equiv_form(1)
e.triangle_in_hom(3) e.triangle_in_hom(4) e.triangle_right
e.triangle_right' e.triangle_right_implies_left)
also have
"... = (\<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w', e, d] \<cdot>\<^sub>D \<a>\<^sub>D[F g \<star>\<^sub>D w', e, d]) \<cdot>\<^sub>D ((\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d)"
using D.comp_assoc by simp
also have "... = (((F g \<star>\<^sub>D w') \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D ((\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d)"
using w' e e.antipar \<beta> D.comp_assoc_assoc' by simp
also have "... = (\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d"
proof -
have "\<guillemotleft>(\<beta> \<star>\<^sub>D e) \<star>\<^sub>D d : ((F g \<star>\<^sub>D w) \<star>\<^sub>D e) \<star>\<^sub>D d \<Rightarrow>\<^sub>D ((F g \<star>\<^sub>D w') \<star>\<^sub>D e) \<star>\<^sub>D d\<guillemotright>"
using w e e.antipar \<beta>
by (intro D.hcomp_in_vhom, auto)
thus ?thesis
using D.comp_cod_arr by auto
qed
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = (F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D ((F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e \<star>\<^sub>D d]) \<cdot>\<^sub>D
(\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e \<star>\<^sub>D d] \<cdot>\<^sub>D (F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (\<a>\<^sub>D[F g, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w', e, d] =
\<a>\<^sub>D[F g, w', e \<star>\<^sub>D d]"
proof -
have "D.seq (F g \<star>\<^sub>D \<a>\<^sub>D[w', e, d])
(\<a>\<^sub>D[F g, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (\<a>\<^sub>D[F g, w', e] \<star>\<^sub>D d))"
using w w' e e.antipar by simp
thus ?thesis
using w w' e e.antipar D.pentagon [of "F g" w' e d] D.invert_side_of_triangle(2)
D.assoc'_eq_inv_assoc D.comp_assoc D.ide_hcomp D.ideD(1)
D.iso_assoc D.hcomp_simps(1) T'.ide_leg1 T.leg1_simps(2-3)
T.tab_simps(2) \<open>src\<^sub>D w' = a\<close> \<open>trg\<^sub>D e = a\<close> \<open>trg\<^sub>D w' = map\<^sub>0 (src\<^sub>C \<rho>)\<close>
e.ide_left e.ide_right preserves_src
by metis
qed
moreover have
"\<a>\<^sub>D[F g \<star>\<^sub>D w, e, d] \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D
(F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) =
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e \<star>\<^sub>D d]"
proof -
have "D.seq (\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] \<star>\<^sub>D d)
(\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]))"
using w w' e e.antipar by simp
moreover have "D.inv \<a>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w, e, d] = \<a>\<^sub>D[F g \<star>\<^sub>D w, e, d]"
using w w' e e.antipar by simp
ultimately show ?thesis
using w w' e e.antipar D.pentagon' [of "F g" w e d]
D.iso_inv_iso D.inv_inv D.comp_assoc D.invert_side_of_triangle(1)
by (metis D.assoc'_simps(3) D.comp_null(2) D.ide_hcomp D.ideD(1)
D.iso_assoc' D.not_arr_null D.seq_if_composable D.src_hcomp T'.ide_leg1
\<open>trg\<^sub>D e = a\<close> a_def e.ide_left e.ide_right)
qed
ultimately show ?thesis
using w w' e e.antipar \<beta> D.comp_assoc by metis
qed
also have "... = (F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', trg\<^sub>D e] \<cdot>\<^sub>D
(((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D ((F g \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, trg\<^sub>D e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(F g \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', e \<star>\<^sub>D d] =
\<a>\<^sub>D[F g, w', trg\<^sub>D e] \<cdot>\<^sub>D ((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>)"
using w' e e.antipar D.assoc_naturality [of "F g" w' \<epsilon>] by simp
moreover have "\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e \<star>\<^sub>D d] \<cdot>\<^sub>D (F g \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) =
((F g \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, trg\<^sub>D e]"
using w e e.antipar D.assoc'_naturality [of "F g" w "D.inv \<epsilon>"]
e.counit_is_iso e.counit_in_hom
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((F g \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D \<a>\<^sub>D[F g, w', trg\<^sub>D e]) \<cdot>\<^sub>D
(\<beta> \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, trg\<^sub>D e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]))"
proof -
have "((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D ((F g \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>) =
\<beta> \<star>\<^sub>D trg\<^sub>D e"
proof -
have "((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D ((F g \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>) =
((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D D.inv \<epsilon>)"
using w w' e e.antipar D.interchange [of \<beta> "F g \<star>\<^sub>D w" "e \<star>\<^sub>D d" "D.inv \<epsilon>"]
D.comp_arr_dom D.comp_cod_arr e.counit_is_iso
by (metis D.in_homE \<beta> d.unit_simps(1) d.unit_simps(3))
also have "... = ((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((F g \<star>\<^sub>D w') \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D trg\<^sub>D e)"
using w w' e e.antipar \<beta> D.interchange [of "F g \<star>\<^sub>D w'" \<beta> "D.inv \<epsilon>" "trg\<^sub>D e"]
D.comp_arr_dom D.comp_cod_arr e.counit_is_iso
by auto
also have
"... = (((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((F g \<star>\<^sub>D w') \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D trg\<^sub>D e)"
using D.comp_assoc by simp
also have "... = ((F g \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon> \<cdot>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D trg\<^sub>D e)"
using w' D.whisker_left [of "F g \<star>\<^sub>D w'"] by simp
also have "... = ((F g \<star>\<^sub>D w') \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D trg\<^sub>D e)"
by (simp add: D.comp_arr_inv')
also have "... = \<beta> \<star>\<^sub>D trg\<^sub>D e"
using \<beta> D.comp_cod_arr D.hseqI'
by (metis D.cod_cod D.hcomp_simps(1) D.hcomp_simps(4)
D.in_homE D.trg.preserves_reflects_arr D.vconn_implies_hpar(1)
D.vconn_implies_hpar(2) D.vconn_implies_hpar(3) D.vconn_implies_hpar(4)
\<open>src\<^sub>D w' = a\<close> \<open>trg\<^sub>D e = a\<close> e.counit_in_hom(2) e.counit_simps(5))
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[F g \<star>\<^sub>D w'] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w]"
using w w' D.runit_hcomp D.runit_hcomp [of "F g" w] by simp
also have "... = \<r>\<^sub>D[F g \<star>\<^sub>D w'] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w'] \<cdot>\<^sub>D \<beta>"
using \<beta> D.runit'_naturality
by (metis D.arr_cod D.arr_dom D.cod_dom D.in_homE D.src.preserves_cod
D.src_dom D.src_hcomp \<open>src\<^sub>D w' = a\<close> \<open>trg\<^sub>D e = a\<close>)
also have "... = (\<r>\<^sub>D[F g \<star>\<^sub>D w'] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F g \<star>\<^sub>D w']) \<cdot>\<^sub>D \<beta>"
using D.comp_assoc by simp
also have "... = \<beta>"
using w' \<beta> D.comp_cod_arr D.comp_arr_inv' D.iso_runit by auto
finally show ?thesis by simp
qed
moreover have "\<theta> = \<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D ?\<gamma>)"
proof -
have "\<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D ?\<gamma>) =
\<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D[w']) \<cdot>\<^sub>D (F f \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using w \<theta> \<gamma> D.whisker_left
by (metis D.arrI D.seqE T'.ide_leg0)
also have
"... = (\<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D[w'])) \<cdot>\<^sub>D (F f \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have 1: "\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] =
\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using w w' e we w'e e.antipar P\<gamma>e D.assoc'_naturality [of "F f" \<gamma>e d]
by (metis D.cod_trg D.in_hhomE D.in_homE D.src_cod D.trg.preserves_cod
T'.leg0_simps(2,4-5) T.tab_simps(2) T.leg0_simps(2)
e.triangle_in_hom(4) e.triangle_right' preserves_src)
also have
2: "... = (\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using D.comp_assoc by simp
also have "... = F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
proof -
have "(\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) =
(F f \<star>\<^sub>D (w' \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using 1 2 e.antipar D.isomorphic_implies_ide(2) w\<^sub>C' w'e D.comp_assoc_assoc'
by auto
also have "... = F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
proof -
have "\<guillemotleft>F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d : F f \<star>\<^sub>D (w \<star>\<^sub>D e) \<star>\<^sub>D d \<Rightarrow>\<^sub>D F f \<star>\<^sub>D (w' \<star>\<^sub>D e) \<star>\<^sub>D d\<guillemotright>"
using we 1 2 e.antipar P\<gamma>e by fastforce
thus ?thesis
using D.comp_cod_arr by blast
qed
finally show ?thesis by blast
qed
finally have
"\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]) =
F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = ((\<theta>' \<cdot>\<^sub>D \<r>\<^sub>D[F f \<star>\<^sub>D w']) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', src\<^sub>D w']) \<cdot>\<^sub>D (F f \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D (\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using w' D.runit_hcomp(3) [of "F f" w'] D.comp_assoc by simp
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D src\<^sub>D w') \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', src\<^sub>D w'] \<cdot>\<^sub>D
(F f \<star>\<^sub>D w' \<star>\<^sub>D \<epsilon>)) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using \<theta>' D.runit_naturality [of \<theta>'] D.comp_assoc by fastforce
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D ((\<theta>' \<star>\<^sub>D src\<^sub>D w') \<cdot>\<^sub>D ((F f \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e \<star>\<^sub>D d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using w' D.assoc'_naturality [of "F f" w' \<epsilon>] D.comp_assoc by simp
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e \<star>\<^sub>D d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<theta>' \<star>\<^sub>D src\<^sub>D w') \<cdot>\<^sub>D ((F f \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) = \<theta>' \<star>\<^sub>D \<epsilon>"
using D.interchange D.comp_arr_dom D.comp_cod_arr
by (metis D.in_homE \<open>src\<^sub>D w' = a\<close> \<open>trg\<^sub>D e = a\<close> \<theta>' e.counit_simps(1)
e.counit_simps(3))
also have "... = (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d)"
using \<theta>' D.interchange [of u \<theta>' \<epsilon> "e \<star>\<^sub>D d"] D.comp_arr_dom D.comp_cod_arr
by auto
finally have "(\<theta>' \<star>\<^sub>D src\<^sub>D w') \<cdot>\<^sub>D ((F f \<star>\<^sub>D w') \<star>\<^sub>D \<epsilon>) = (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e \<star>\<^sub>D d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) \<cdot>\<^sub>D ((\<a>\<^sub>D[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d])) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) =
(F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D[F f, w \<star>\<^sub>D e, d]"
using D.assoc_naturality [of "F f" \<gamma>e d]
by (metis D.cod_trg D.in_hhomE D.in_homE D.src_cod D.trg.preserves_cod P\<gamma>e
T'.leg0_simps(2,4-5) T.tab_simps(2) T.leg0_simps(2) e e.antipar(1)
e.triangle_in_hom(4) e.triangle_right' preserves_src w'e)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D
(\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e \<star>\<^sub>D d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<a>\<^sub>D[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) =
F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]"
using w D.comp_cod_arr D.comp_assoc_assoc' by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D
((\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e \<star>\<^sub>D d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D
((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d =
\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d]"
proof -
have "\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] =
\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using P\<gamma>e e.antipar D.assoc'_naturality
by (metis D.in_hhom_def D.in_homE D.vconn_implies_hpar(1)
D.vconn_implies_hpar(2) T'.leg0_simps(2,4-5)
T.leg0_simps(2) T.tab_simps(2) \<open>src\<^sub>D e = map\<^sub>0 a\<^sub>C\<close>
d.triangle_equiv_form(1) d.triangle_in_hom(3) d.triangle_left
preserves_src we)
also have
"... = (\<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using D.comp_assoc by simp
also have "... = (F f \<star>\<^sub>D (w' \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d)"
using w'e D.isomorphic_implies_ide(2) w\<^sub>C' D.comp_assoc_assoc' by auto
also have "... = F f \<star>\<^sub>D \<gamma>e \<star>\<^sub>D d"
using D.comp_cod_arr
by (metis D.comp_cod_arr D.comp_null(2) D.hseq_char D.hseq_char'
D.in_homE D.whisker_left D.whisker_right P\<gamma>e T'.ide_leg0 e.ide_right)
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D[F f \<star>\<^sub>D w', e, d]) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D
((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e \<star>\<^sub>D d]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d] =
\<a>\<^sub>D[F f \<star>\<^sub>D w', e, d] \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<star>\<^sub>D d)"
proof -
have "\<a>\<^sub>D[F f, w', e \<star>\<^sub>D d] \<cdot>\<^sub>D \<a>\<^sub>D[F f \<star>\<^sub>D w', e, d] =
((F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d]) \<cdot>\<^sub>D (\<a>\<^sub>D[F f, w', e] \<star>\<^sub>D d)"
using w' D.pentagon D.comp_assoc by simp
moreover have "D.seq \<a>\<^sub>D[F f, w', e \<star>\<^sub>D d] \<a>\<^sub>D[F f \<star>\<^sub>D w', e, d]"
using w' by simp
moreover have "D.inv (\<a>\<^sub>D[F f, w', e] \<star>\<^sub>D d) = \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<star>\<^sub>D d"
using w' by simp
ultimately show ?thesis
using w' D.comp_assoc
D.invert_opposite_sides_of_square
[of "\<a>\<^sub>D[F f, w', e \<star>\<^sub>D d]" "\<a>\<^sub>D[F f \<star>\<^sub>D w', e, d]"
"(F f \<star>\<^sub>D \<a>\<^sub>D[w', e, d]) \<cdot>\<^sub>D \<a>\<^sub>D[F f, w' \<star>\<^sub>D e, d]"
"\<a>\<^sub>D[F f, w', e] \<star>\<^sub>D d"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have
"... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D
(((\<theta>' \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<theta>' \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D[F f \<star>\<^sub>D w', e, d] = \<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D ((\<theta>' \<star>\<^sub>D e) \<star>\<^sub>D d)"
using w' \<theta>' e.ide_left e.ide_right e.antipar D.assoc_naturality [of \<theta>' e d]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D
((\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "((\<theta>' \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<star>\<^sub>D d) \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d) =
(\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e) \<star>\<^sub>D d"
using w' w'e \<theta>' \<theta>\<^sub>C e.ide_left e.ide_right e.antipar D.whisker_right
by (metis (full_types) C.arrI D.cod_comp D.seqE D.seqI F\<theta>\<^sub>C_def P\<gamma>e
preserves_arr)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D
((\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e) =
\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using P\<gamma>e by simp
moreover have "D.arr (\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e))"
by (metis C.in_homE D.comp_assoc D.comp_null(1) D.ext F\<theta>\<^sub>C_def P\<gamma>e \<theta>\<^sub>C
preserves_arr)
moreover have "D.arr (\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e])"
using P\<gamma>e calculation(2) by auto
ultimately have "(\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>e) =
(\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using \<psi> \<theta>\<^sub>C F\<theta>\<^sub>C_def D.iso_is_section D.section_is_mono
by (metis D.monoE)
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D
((\<theta> \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D ((\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<star>\<^sub>D d) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w \<star>\<^sub>D e, d] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d])) \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<star>\<^sub>D d =
((\<theta> \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] \<star>\<^sub>D d)"
proof -
have "D.arr ((\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e])"
by (metis C.arrI D.cod_comp D.seqE D.seqI F\<theta>\<^sub>C_def \<theta>\<^sub>C preserves_arr)
thus ?thesis
using D.whisker_right e.ide_right by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D
(((\<theta> \<star>\<^sub>D e) \<star>\<^sub>D d) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f \<star>\<^sub>D w, e, d]) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e \<star>\<^sub>D d] \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using w D.pentagon' D.comp_assoc by simp
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((\<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[u, e, d]) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e \<star>\<^sub>D d] \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using \<theta> e.antipar D.assoc'_naturality [of \<theta> e d] D.comp_assoc by fastforce
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D (\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e \<star>\<^sub>D d] \<cdot>\<^sub>D
(F f \<star>\<^sub>D w \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(\<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u, e, d]) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d) = \<theta> \<star>\<^sub>D e \<star>\<^sub>D d"
proof -
have "(\<a>\<^sub>D[u, e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[u, e, d]) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d) =
(u \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d)"
using \<theta> ue e.ide_left e.ide_right e.antipar D.comp_arr_inv' D.comp_cod_arr
by auto
also have "... = \<theta> \<star>\<^sub>D e \<star>\<^sub>D d"
using ue e.ide_left e.ide_right e.antipar D.hcomp_simps(4) D.hseq_char' \<theta>
D.comp_cod_arr [of "\<theta> \<star>\<^sub>D e \<star>\<^sub>D d" "u \<star>\<^sub>D e \<star>\<^sub>D d"]
by force
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D ((u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d)) \<cdot>\<^sub>D ((F f \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
using w e.antipar D.assoc'_naturality [of "F f" w "D.inv \<epsilon>"] D.comp_assoc by simp
also have
"... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D (((F f \<star>\<^sub>D w) \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((F f \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e]) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(u \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e \<star>\<^sub>D d) = (\<theta> \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D ((F f \<star>\<^sub>D w) \<star>\<^sub>D \<epsilon>)"
using \<theta> e.antipar D.interchange D.comp_arr_dom D.comp_cod_arr
by (metis D.in_homE \<open>trg\<^sub>D e = a\<close> e.counit_simps(1-3,5))
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[u] \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w])"
proof -
have "(((F f \<star>\<^sub>D w) \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((F f \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e] =
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e]"
proof -
have "(((F f \<star>\<^sub>D w) \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D ((F f \<star>\<^sub>D w) \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e] =
((F f \<star>\<^sub>D w) \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e]"
using w e.ide_left e.ide_right e.antipar e.counit_is_iso D.comp_arr_inv'
D.comp_assoc D.whisker_left
by (metis D.ide_hcomp D.seqI' T'.ide_leg0 T'.leg1_simps(3)
T.leg1_simps(2-3) T.tab_simps(2) \<open>trg\<^sub>D w = map\<^sub>0 (src\<^sub>C \<rho>)\<close>
d.unit_in_vhom e.counit_in_hom(2) e.counit_simps(3) preserves_src)
also have "... = \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, trg\<^sub>D e]"
using w D.comp_cod_arr D.assoc'_in_hom(2) [of "F f" w "trg\<^sub>D e"]
\<open>trg\<^sub>D e = a\<close> \<open>trg\<^sub>D w = map\<^sub>0 (src\<^sub>C \<rho>)\<close>
by (metis D.assoc'_is_natural_1 D.ideD(1) D.ideD(2) D.trg.preserves_ide
D.trg_trg T'.leg0_simps(2,4) T'.leg1_simps(3)
T.leg1_simps(2-3) T.tab_simps(2) a_def e.ide_left
preserves_src)
finally show ?thesis by blast
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (\<r>\<^sub>D[u] \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D trg\<^sub>D e)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f \<star>\<^sub>D w]"
using w D.runit_hcomp(2) [of "F f" w] D.comp_assoc by simp
also have 1: "... = (\<theta> \<cdot>\<^sub>D \<r>\<^sub>D[F f \<star>\<^sub>D w]) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f \<star>\<^sub>D w]"
using \<theta> D.runit_naturality [of \<theta>] by auto
also have "... = \<theta>"
using w \<theta> D.comp_arr_dom D.comp_assoc
by (metis D.hcomp_arr_obj(2) D.in_homE D.obj_src 1 \<open>src\<^sub>D \<theta> = a\<close> \<open>trg\<^sub>D e = a\<close>)
finally show ?thesis by simp
qed
ultimately show "\<guillemotleft>?\<gamma> : w \<Rightarrow>\<^sub>D w'\<guillemotright> \<and> \<beta> = F g \<star>\<^sub>D ?\<gamma> \<and> \<theta> = \<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D ?\<gamma>)"
by simp
show "\<And>\<gamma>'. \<guillemotleft>\<gamma>' : w \<Rightarrow>\<^sub>D w'\<guillemotright> \<and> \<beta> = F g \<star>\<^sub>D \<gamma>' \<and> \<theta> = \<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>') \<Longrightarrow> \<gamma>' = ?\<gamma>"
proof -
fix \<gamma>'
assume \<gamma>': "\<guillemotleft>\<gamma>' : w \<Rightarrow>\<^sub>D w'\<guillemotright> \<and> \<beta> = F g \<star>\<^sub>D \<gamma>' \<and> \<theta> = \<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>')"
show "\<gamma>' = ?\<gamma>"
proof -
have "?\<gamma> = \<r>\<^sub>D[w'] \<cdot>\<^sub>D (w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<a>\<^sub>D[w', e, d] \<cdot>\<^sub>D ((\<gamma>' \<star>\<^sub>D e) \<star>\<^sub>D d)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
proof -
have "\<gamma>e = \<gamma>' \<star>\<^sub>D e"
proof -
have "\<guillemotleft>\<gamma>' \<star>\<^sub>D e : w \<star>\<^sub>D e \<Rightarrow>\<^sub>D w' \<star>\<^sub>D e\<guillemotright>"
using \<gamma>' by (intro D.hcomp_in_vhom, auto)
moreover have
"\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] = F g \<star>\<^sub>D \<gamma>' \<star>\<^sub>D e"
proof -
have "\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D (\<beta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e] =
\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D ((F g \<star>\<^sub>D \<gamma>') \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w, e]"
using \<gamma>' by simp
also have "... = \<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w', e] \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>' \<star>\<^sub>D e)"
using \<gamma>' D.assoc_naturality
by (metis D.assoc'_naturality D.hcomp_in_vhomE D.ideD(2) D.ideD(3)
D.in_homE T'.leg1_simps(5-6) \<beta>
\<open>\<guillemotleft>\<gamma>' \<star>\<^sub>D e : w \<star>\<^sub>D e \<Rightarrow>\<^sub>D w' \<star>\<^sub>D e\<guillemotright>\<close> e.ide_left)
also have "... = (\<a>\<^sub>D[F g, w', e] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F g, w', e]) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>' \<star>\<^sub>D e)"
using D.comp_assoc by simp
also have "... = F g \<star>\<^sub>D \<gamma>' \<star>\<^sub>D e"
by (metis D.hcomp_reassoc(2) D.in_homE D.not_arr_null D.seq_if_composable
T'.leg1_simps(2,5-6) \<beta> \<gamma>' calculation
\<open>\<guillemotleft>\<gamma>' \<star>\<^sub>D e : w \<star>\<^sub>D e \<Rightarrow>\<^sub>D w' \<star>\<^sub>D e\<guillemotright>\<close> e.triangle_equiv_form(1)
e.triangle_in_hom(3) e.triangle_right e.triangle_right_implies_left)
finally show ?thesis by simp
qed
moreover have "\<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e] =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>' \<star>\<^sub>D e)"
proof -
have "\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w', e] \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>' \<star>\<^sub>D e) =
\<psi> \<cdot>\<^sub>D (\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>') \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using \<gamma>' \<theta> e.ide_left D.assoc'_naturality
by (metis D.hcomp_in_vhomE D.ideD(2) D.ideD(3) D.in_homE
T'.leg0_simps(2,4-5) T'.leg1_simps(3) \<beta> calculation(1))
also have "... = \<psi> \<cdot>\<^sub>D ((\<theta>' \<star>\<^sub>D e) \<cdot>\<^sub>D ((F f \<star>\<^sub>D \<gamma>') \<star>\<^sub>D e)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using D.comp_assoc by simp
also have "... = \<psi> \<cdot>\<^sub>D (\<theta>' \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>') \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using D.whisker_right \<gamma>' \<theta> by auto
also have "... = \<psi> \<cdot>\<^sub>D (\<theta> \<star>\<^sub>D e) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, w, e]"
using \<gamma>' by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using UN by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<r>\<^sub>D[w'] \<cdot>\<^sub>D ((w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<gamma>' \<star>\<^sub>D e \<star>\<^sub>D d)) \<cdot>\<^sub>D \<a>\<^sub>D[w, e, d] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using w' \<gamma>' D.comp_assoc D.assoc_naturality
by (metis D.in_homE D.src_dom \<open>trg\<^sub>D e = a\<close> a_def e.antipar(1)
e.triangle_equiv_form(1) e.triangle_in_hom(3-4)
e.triangle_right e.triangle_right' e.triangle_right_implies_left)
also have "... = (\<r>\<^sub>D[w'] \<cdot>\<^sub>D (\<gamma>' \<star>\<^sub>D trg\<^sub>D e)) \<cdot>\<^sub>D (w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[w, e, d] \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
proof -
have "(w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<gamma>' \<star>\<^sub>D e \<star>\<^sub>D d) = \<gamma>' \<star>\<^sub>D \<epsilon>"
using w' \<gamma>' e.antipar D.comp_arr_dom D.comp_cod_arr
D.interchange [of w' \<gamma>' \<epsilon> "e \<star>\<^sub>D d"]
by auto
also have "... = (\<gamma>' \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D (w \<star>\<^sub>D \<epsilon>)"
using w \<gamma>' e.antipar D.comp_arr_dom D.comp_cod_arr D.interchange
by (metis D.in_homE \<open>trg\<^sub>D e = a\<close> e.counit_simps(1) e.counit_simps(3,5))
finally have "(w' \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<gamma>' \<star>\<^sub>D e \<star>\<^sub>D d) = (\<gamma>' \<star>\<^sub>D trg\<^sub>D e) \<cdot>\<^sub>D (w \<star>\<^sub>D \<epsilon>)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<gamma>' \<cdot>\<^sub>D \<r>\<^sub>D[w] \<cdot>\<^sub>D (w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[w, e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D
(w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using \<gamma>' D.runit_naturality D.comp_assoc
by (metis D.in_homE D.src_dom \<open>trg\<^sub>D e = a\<close> a_def)
also have "... = \<gamma>'"
proof -
have "\<r>\<^sub>D[w] \<cdot>\<^sub>D (w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[w, e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[w] =
\<r>\<^sub>D[w] \<cdot>\<^sub>D ((w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (\<a>\<^sub>D[w, e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d]) \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using D.comp_assoc by simp
also have "... = \<r>\<^sub>D[w] \<cdot>\<^sub>D ((w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (w \<star>\<^sub>D e \<star>\<^sub>D d) \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D
\<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using w \<gamma> e.ide_left e.ide_right we e.antipar D.comp_assoc_assoc'(1)
\<open>trg\<^sub>D e = a\<close> a_def
by presburger
also have "... = \<r>\<^sub>D[w] \<cdot>\<^sub>D ((w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D (w \<star>\<^sub>D D.inv \<epsilon>)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using w \<gamma> e.ide_left e.ide_right we e.antipar D.comp_cod_arr
by (metis D.whisker_left d.unit_simps(1,3))
also have "... = \<r>\<^sub>D[w] \<cdot>\<^sub>D (w \<star>\<^sub>D src\<^sub>D w) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using w e.counit_is_iso C.comp_arr_inv'
by (metis D.comp_arr_inv' D.seqI' D.whisker_left \<open>trg\<^sub>D e = a\<close> a_def
d.unit_in_vhom e.counit_in_hom(2) e.counit_simps(3))
also have "... = \<r>\<^sub>D[w] \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w]"
using w e.antipar D.comp_cod_arr by simp
also have "... = w"
using w
by (simp add: D.comp_arr_inv')
finally have "\<r>\<^sub>D[w] \<cdot>\<^sub>D (w \<star>\<^sub>D \<epsilon>) \<cdot>\<^sub>D \<a>\<^sub>D[w, e, d] \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[w, e, d] \<cdot>\<^sub>D
(w \<star>\<^sub>D D.inv \<epsilon>) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[w] = w"
by simp
thus ?thesis
using \<gamma>' D.comp_arr_dom by auto
qed
finally show ?thesis by simp
qed
qed
qed
qed
qed
show ?thesis ..
qed
lemma reflects_tabulation:
assumes "C.ide r" and "C.ide f" and "\<guillemotleft>\<rho> : g \<Rightarrow>\<^sub>C r \<star>\<^sub>C f\<guillemotright>"
assumes "tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F r) (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>) (F f) (F g)"
shows "tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> f g"
proof -
interpret \<rho>': tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>F r\<close> \<open>D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>\<close> \<open>F f\<close> \<open>F g\<close>
using assms by auto
interpret \<rho>: tabulation_data V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> f g
using assms by (unfold_locales, simp_all)
interpret \<rho>: tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> f g
proof
show "\<And>u \<omega>. \<lbrakk> C.ide u; \<guillemotleft>\<omega> : C.dom \<omega> \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. C.ide w \<and> \<guillemotleft>\<theta> : f \<star>\<^sub>C w \<Rightarrow>\<^sub>C u\<guillemotright> \<and> \<guillemotleft>\<nu> : C.dom \<omega> \<Rightarrow>\<^sub>C g \<star>\<^sub>C w\<guillemotright> \<and>
C.iso \<nu> \<and> \<rho>.composite_cell w \<theta> \<cdot>\<^sub>C \<nu> = \<omega>"
proof -
fix u \<omega>
assume u: "C.ide u"
assume \<omega>: "\<guillemotleft>\<omega> : C.dom \<omega> \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<guillemotright>"
have hseq_ru: "src\<^sub>C r = trg\<^sub>C u"
using \<omega> C.ide_cod C.ideD(1) by fastforce
hence 1: "\<guillemotleft>D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F \<omega> : F (C.dom \<omega>) \<Rightarrow>\<^sub>D F r \<star>\<^sub>D F u\<guillemotright>"
using assms u \<omega> cmp_in_hom cmp_components_are_iso
by (intro D.comp_in_homI, auto)
hence 2: "D.dom (D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F \<omega>) = F (C.dom \<omega>)"
by auto
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "D.ide w \<and> \<guillemotleft>\<theta> : F f \<star>\<^sub>D w \<Rightarrow>\<^sub>D F u\<guillemotright> \<and>
\<guillemotleft>\<nu> : F (C.dom \<omega>) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w\<guillemotright> \<and> D.iso \<nu> \<and>
\<rho>'.composite_cell w \<theta> \<cdot>\<^sub>D \<nu> = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F \<omega>"
using 1 2 u \<rho>'.T1 [of "F u" "D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F \<omega>"] by auto
have hseq_Ff_w: "src\<^sub>D (F f) = trg\<^sub>D w"
using u \<omega> w\<theta>\<nu>
by (metis "1" D.arrI D.not_arr_null D.seqE D.seq_if_composable \<rho>'.tab_simps(2))
have hseq_Fg_w: "src\<^sub>D (F g) = trg\<^sub>D w"
using u \<omega> w\<theta>\<nu> by (simp add: hseq_Ff_w)
have w: "\<guillemotleft>w : map\<^sub>0 (src\<^sub>C \<omega>) \<rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C f)\<guillemotright>"
using u \<omega> w\<theta>\<nu> hseq_Fg_w
by (metis "1" C.arrI D.arrI D.hseqI' D.ideD(1) D.in_hhom_def D.src_hcomp
D.src_vcomp D.vconn_implies_hpar(1) D.vconn_implies_hpar(3)
D.vseq_implies_hpar(1) \<rho>'.leg1_simps(2) \<rho>.leg0_simps(2) hseq_Ff_w
preserves_src)
obtain w' where w': "\<guillemotleft>w' : src\<^sub>C \<omega> \<rightarrow>\<^sub>C src\<^sub>C f\<guillemotright> \<and> C.ide w' \<and> D.isomorphic (F w') w"
using assms w \<omega> w\<theta>\<nu> locally_essentially_surjective by force
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : F w' \<Rightarrow>\<^sub>D w\<guillemotright> \<and> D.iso \<phi>"
using w' D.isomorphic_def by blast
have src_fw': "src\<^sub>C (f \<star>\<^sub>C w') = src\<^sub>C u"
using u w' \<omega>
by (metis C.hseqI' C.ideD(1) C.in_hhomE C.src_hcomp C.vconn_implies_hpar(1)
C.vconn_implies_hpar(3) \<rho>.base_simps(2) \<rho>.leg0_in_hom(1) hseq_ru)
have 3: "\<guillemotleft>\<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w')) : F (f \<star>\<^sub>C w') \<Rightarrow>\<^sub>D F u\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>D.inv (\<Phi> (f, w')) : F (f \<star>\<^sub>C w') \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F w'\<guillemotright>"
using assms w' cmp_in_hom cmp_components_are_iso by auto
show "\<guillemotleft>F f \<star>\<^sub>D \<phi> : F f \<star>\<^sub>D F w' \<Rightarrow>\<^sub>D F f \<star>\<^sub>D w\<guillemotright>"
using \<phi> \<rho>'.leg0_in_hom(2) w' by fastforce
show "\<guillemotleft>\<theta> : F f \<star>\<^sub>D w \<Rightarrow>\<^sub>D F u\<guillemotright>"
using w\<theta>\<nu> by simp
qed
have 4: "\<exists>\<theta>'. \<guillemotleft>\<theta>' : f \<star>\<^sub>C w' \<Rightarrow>\<^sub>C u\<guillemotright> \<and> F \<theta>' = \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w'))"
using w' u hseq_ru src_fw' 3 locally_full by auto
obtain \<theta>' where
\<theta>': "\<guillemotleft>\<theta>' : f \<star>\<^sub>C w' \<Rightarrow>\<^sub>C u\<guillemotright> \<and> F \<theta>' = \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w'))"
using 4 by auto
have 5: "\<guillemotleft>\<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D \<nu> : F (C.dom \<omega>) \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C w')\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<nu> : F (C.dom \<omega>) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D w\<guillemotright>"
using w\<theta>\<nu> by simp
show "\<guillemotleft>F g \<star>\<^sub>D D.inv \<phi> : F g \<star>\<^sub>D w \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F w'\<guillemotright>"
using assms \<phi>
by (meson D.hcomp_in_vhom D.inv_in_hom \<rho>'.leg1_in_hom(2) hseq_Fg_w)
show "\<guillemotleft>\<Phi> (g, w') : F g \<star>\<^sub>D F w' \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C w')\<guillemotright>"
using assms w' cmp_in_hom by auto
qed
have 6: "\<exists>\<nu>'. \<guillemotleft>\<nu>' : C.dom \<omega> \<Rightarrow>\<^sub>C g \<star>\<^sub>C w'\<guillemotright> \<and>
F \<nu>' = \<Phi>(g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D \<nu>"
using u w' \<omega> C.in_hhom_def hseq_ru C.hseqI' C.hcomp_simps(1-2)
by (metis "5" C.arrI C.ide_hcomp C.ideD(1) C.ide_dom C.vconn_implies_hpar(1,4)
\<rho>.base_simps(2) \<rho>.ide_leg1 \<rho>.leg1_in_hom(1) locally_full)
obtain \<nu>' where
\<nu>': "\<guillemotleft>\<nu>' : C.dom \<omega> \<Rightarrow>\<^sub>C g \<star>\<^sub>C w'\<guillemotright> \<and> F \<nu>' = \<Phi>(g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D \<nu>"
using 6 by auto
have "C.ide w' \<and> \<guillemotleft>\<theta>' : f \<star>\<^sub>C w' \<Rightarrow>\<^sub>C u\<guillemotright> \<and> \<guillemotleft>\<nu>' : C.dom \<omega> \<Rightarrow>\<^sub>C g \<star>\<^sub>C w'\<guillemotright> \<and> C.iso \<nu>' \<and>
\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>' = \<omega>"
using w' \<theta>' \<nu>'
apply (intro conjI)
apply auto
proof -
show "C.iso \<nu>'"
proof -
have "D.iso (F \<nu>')"
proof -
have "D.iso (\<Phi>(g, w'))"
using w' cmp_components_are_iso by auto
moreover have "D.iso (F g \<star>\<^sub>D D.inv \<phi>)"
using \<phi>
by (meson "5" D.arrI D.iso_hcomp D.hseq_char' D.ide_is_iso D.iso_inv_iso
D.seqE D.seq_if_composable \<rho>'.ide_leg1)
moreover have "D.iso \<nu>"
using w\<theta>\<nu> by simp
ultimately show ?thesis
using \<nu>' D.isos_compose
by (metis "5" D.arrI D.seqE)
qed
thus ?thesis using reflects_iso by blast
qed
have 7: "\<guillemotleft>\<rho>.composite_cell w' \<theta>' : g \<star>\<^sub>C w' \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<guillemotright>"
using u w' \<theta>' \<rho>.composite_cell_in_hom hseq_ru src_fw' C.hseqI'
by (metis C.in_hhomE C.hcomp_simps(1) \<rho>.leg0_simps(2))
hence 8: "\<guillemotleft>\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>' : C.dom \<omega> \<Rightarrow>\<^sub>C r \<star>\<^sub>C u\<guillemotright>"
using \<nu>' by blast
show "\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>' = \<omega>"
proof -
have 1: "C.par (\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>') \<omega>"
using \<omega> 8 hseq_ru C.hseqI' C.in_homE by metis
moreover have "F (\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>') = F \<omega>"
proof -
have "F (\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>') =
F (r \<star>\<^sub>C \<theta>') \<cdot>\<^sub>D F \<a>\<^sub>C[r, f, w'] \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w') \<cdot>\<^sub>D F \<nu>'"
using w' \<theta>' \<nu>' 1 C.comp_assoc
by (metis C.seqE preserves_comp)
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D
\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w'))) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D
((D.inv (\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D
\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w')) \<cdot>\<^sub>D D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<nu>'"
proof -
have "F \<a>\<^sub>C[r, f, w'] =
\<Phi> (r, f \<star>\<^sub>C w') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w'))"
using assms w'
by (simp add: C.in_hhom_def preserves_assoc(1))
moreover have
"F (r \<star>\<^sub>C \<theta>') = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D D.inv (\<Phi> (r, f \<star>\<^sub>C w'))"
using assms \<theta>' preserves_hcomp [of r \<theta>']
by (metis "1" C.in_homE C.seqE \<rho>.base_simps(3) \<rho>.base_simps(4))
moreover have
"F (\<rho> \<star>\<^sub>C w') = \<Phi> (r \<star>\<^sub>C f, w') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D D.inv (\<Phi> (g, w'))"
using w' preserves_hcomp [of \<rho> w'] by auto
ultimately show ?thesis
by (simp add: D.comp_assoc)
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D
(F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<nu>'"
proof -
have "(D.inv (\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) =
F r \<star>\<^sub>D \<Phi> (f, w')"
using w' cmp_components_are_iso D.comp_cod_arr C.hseqI' D.hseqI'
C.in_hhom_def C.trg_hcomp D.comp_inv_arr' C.ide_hcomp
by (metis C.ideD(1) D.hcomp_simps(4) cmp_simps(1,3-5)
\<rho>'.leg0_simps(3) \<rho>'.base_simps(2,4) \<rho>.ide_leg0 \<rho>.ide_base
\<rho>.leg0_simps(3))
moreover have "(D.inv (\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D \<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w') =
F \<rho> \<star>\<^sub>D F w'"
using w' D.comp_inv_arr' hseq_Fg_w D.comp_cod_arr by auto
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w'))) \<cdot>\<^sub>D
(F r \<star>\<^sub>D \<Phi> (f, w'))) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D
((D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w')) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D \<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D \<nu>"
using w' \<theta>' \<nu>' D.comp_assoc by simp
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w')) \<cdot>\<^sub>D
\<Phi> (f, w')) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D
F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D ((D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D \<Phi> (g, w')) \<cdot>\<^sub>D
(F g \<star>\<^sub>D D.inv \<phi>)) \<cdot>\<^sub>D \<nu>"
proof -
have "(F r \<star>\<^sub>D \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w'))) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) =
F r \<star>\<^sub>D (\<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w'))) \<cdot>\<^sub>D \<Phi> (f, w')"
proof -
have "D.seq (\<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w'))) (\<Phi> (f, w'))"
using assms 3 \<rho>.ide_base w' w\<theta>\<nu> cmp_in_hom [of f w'] cmp_components_are_iso
C.in_hhom_def
apply (intro D.seqI)
using C.in_hhom_def
apply auto[3]
apply blast
by auto
thus ?thesis
using assms w' w\<theta>\<nu> cmp_in_hom cmp_components_are_iso D.whisker_left
by simp
qed
moreover have "(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w') =
D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w'"
using w' D.whisker_right by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>)) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D
(F g \<star>\<^sub>D D.inv \<phi>)) \<cdot>\<^sub>D \<nu>"
proof -
have "(F f \<star>\<^sub>D \<phi>) \<cdot>\<^sub>D D.inv (\<Phi> (f, w')) \<cdot>\<^sub>D \<Phi> (f, w') = F f \<star>\<^sub>D \<phi>"
using assms(2) w' \<phi> 3 cmp_components_are_iso cmp_in_hom D.hseqI' D.comp_inv_arr'
D.comp_arr_dom
by (metis C.in_hhom_def D.arrI D.cod_inv D.seqE)
moreover have "(D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D \<Phi> (g, w')) \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) =
F g \<star>\<^sub>D D.inv \<phi>"
using assms w' \<phi> 3 cmp_components_are_iso cmp_in_hom D.hseqI'
D.comp_inv_arr' D.comp_cod_arr
by (metis "5" C.in_hhom_def D.arrI D.comp_assoc D.seqE \<rho>.ide_leg1
\<rho>.leg1_simps(3))
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>)) \<cdot>\<^sub>D
(\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F f) \<star>\<^sub>D D.inv \<phi>)) \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D w) \<cdot>\<^sub>D \<nu>"
proof -
have "(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) =
D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D D.inv \<phi>"
using assms w' \<phi> cmp_in_hom cmp_components_are_iso D.comp_arr_dom
D.comp_cod_arr
D.interchange [of "D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>" "F g" "F w'" "D.inv \<phi>"]
by auto
also have "... = ((F r \<star>\<^sub>D F f) \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D w)"
using assms w' \<phi> cmp_components_are_iso D.comp_arr_dom D.comp_cod_arr
D.interchange [of "F r \<star>\<^sub>D F f" "D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho>" "D.inv \<phi>" w]
by auto
finally have "(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D (F g \<star>\<^sub>D D.inv \<phi>) =
((F r \<star>\<^sub>D F f) \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D w)"
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D ((F r \<star>\<^sub>D \<theta> \<cdot>\<^sub>D (F f \<star>\<^sub>D \<phi>)) \<cdot>\<^sub>D
(F r \<star>\<^sub>D F f \<star>\<^sub>D D.inv \<phi>)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D w) \<cdot>\<^sub>D \<nu>"
proof -
have "\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D ((F r \<star>\<^sub>D F f) \<star>\<^sub>D D.inv \<phi>) =
(F r \<star>\<^sub>D F f \<star>\<^sub>D D.inv \<phi>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w]"
proof -
have "src\<^sub>D (F r) = trg\<^sub>D (F f)"
by simp
moreover have "src\<^sub>D (F f) = trg\<^sub>D (D.inv \<phi>)"
using \<phi>
by (metis "5" D.arrI D.hseqE D.seqE \<rho>'.leg1_simps(3))
ultimately show ?thesis
using assms w' \<phi> D.assoc_naturality [of "F r" "F f" "D.inv \<phi>"] by auto
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<theta>) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D w) \<cdot>\<^sub>D \<nu>"
using assms \<phi> w\<theta>\<nu> D.comp_arr_inv' D.comp_arr_dom D.comp_cod_arr
D.whisker_left D.whisker_left D.comp_assoc
by (metis D.ideD(1) D.in_homE \<rho>'.ide_base tabulation_data.leg0_simps(1)
tabulation_def)
also have "... = (\<Phi> (r, u) \<cdot>\<^sub>D D.inv (\<Phi> (r, u))) \<cdot>\<^sub>D F \<omega>"
using w\<theta>\<nu> D.comp_assoc by simp
also have "... = F \<omega>"
using u \<omega> cmp_in_hom D.comp_arr_inv'
by (metis C.in_homE cmp_components_are_iso cmp_simps(5) \<rho>.ide_base
as_nat_trans.is_natural_1 as_nat_trans.naturality hseq_ru)
finally show ?thesis by blast
qed
ultimately show ?thesis
using is_faithful [of "\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<nu>'" \<omega>] by simp
qed
qed
thus "\<exists>w \<theta> \<nu>. C.ide w \<and> \<guillemotleft>\<theta> : f \<star>\<^sub>C w \<Rightarrow>\<^sub>C u\<guillemotright> \<and> \<guillemotleft>\<nu> : C.dom \<omega> \<Rightarrow>\<^sub>C g \<star>\<^sub>C w\<guillemotright> \<and>
C.iso \<nu> \<and> \<rho>.composite_cell w \<theta> \<cdot>\<^sub>C \<nu> = \<omega>"
by auto
qed
show "\<And>u w w' \<theta> \<theta>' \<beta>. \<lbrakk> C.ide w; C.ide w'; \<guillemotleft>\<theta> : f \<star>\<^sub>C w \<Rightarrow>\<^sub>C u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star>\<^sub>C w' \<Rightarrow>\<^sub>C u\<guillemotright>;
\<guillemotleft>\<beta> : g \<star>\<^sub>C w \<Rightarrow>\<^sub>C g \<star>\<^sub>C w'\<guillemotright>;
\<rho>.composite_cell w \<theta> = \<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<beta> \<rbrakk>
\<Longrightarrow> \<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma> \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume w: "C.ide w"
assume w': "C.ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : f \<star>\<^sub>C w \<Rightarrow>\<^sub>C u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : f \<star>\<^sub>C w' \<Rightarrow>\<^sub>C u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : g \<star>\<^sub>C w \<Rightarrow>\<^sub>C g \<star>\<^sub>C w'\<guillemotright>"
assume eq: "\<rho>.composite_cell w \<theta> = \<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<beta>"
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma> \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)"
proof -
have hseq_ru: "src\<^sub>C r = trg\<^sub>C u"
using w \<theta>
by (metis C.hseq_char' C.in_homE C.trg.is_extensional C.trg.preserves_hom
C.trg_hcomp C.vconn_implies_hpar(2) C.vconn_implies_hpar(4) \<rho>.leg0_simps(3))
have hseq_fw: "src\<^sub>C f = trg\<^sub>C w \<and> src\<^sub>C f = trg\<^sub>C w'"
using w w' \<rho>.ide_leg0 \<theta> \<theta>'
by (metis C.horizontal_homs_axioms C.ideD(1) C.in_homE C.not_arr_null
C.seq_if_composable category.ide_dom horizontal_homs_def)
have hseq_gw: "src\<^sub>C g = trg\<^sub>C w \<and> src\<^sub>C g = trg\<^sub>C w'"
using w w' \<rho>.ide_leg0 \<theta> \<theta>' \<open>src\<^sub>C f = trg\<^sub>C w \<and> src\<^sub>C f = trg\<^sub>C w'\<close> by auto
have *: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : F w \<Rightarrow>\<^sub>D F w'\<guillemotright> \<and>
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) = F g \<star>\<^sub>D \<gamma> \<and>
F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) = (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>)"
proof -
have "D.ide (F w) \<and> D.ide (F w')"
using w w' by simp
moreover have 1: "\<guillemotleft>F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) : F f \<star>\<^sub>D F w \<Rightarrow>\<^sub>D F u\<guillemotright>"
using w \<theta> cmp_in_hom \<rho>.ide_leg0 hseq_fw by blast
moreover have 2: "\<guillemotleft>F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') : F f \<star>\<^sub>D F w' \<Rightarrow>\<^sub>D F u\<guillemotright>"
using w' \<theta>' cmp_in_hom \<rho>.ide_leg0 hseq_fw by blast
moreover have
"\<guillemotleft>D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) : F g \<star>\<^sub>D F w \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F w'\<guillemotright>"
using w w' \<beta> \<rho>.ide_leg1 cmp_in_hom cmp_components_are_iso hseq_gw preserves_hom
by fastforce
moreover have "\<rho>'.composite_cell (F w) (F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)) =
\<rho>'.composite_cell (F w') (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
proof -
have "\<rho>'.composite_cell (F w') (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) =
(F r \<star>\<^sub>D F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<cdot>\<^sub>D F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
using D.comp_assoc by simp
also have "... =
(F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
using w' \<theta>' 2 D.whisker_left D.whisker_right D.comp_assoc by auto
also have "... = (F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D
\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w'))) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D
((D.inv (\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D
\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w')) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
proof -
have "(D.inv (\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) =
F r \<star>\<^sub>D \<Phi> (f, w')"
using w' cmp_components_are_iso D.comp_cod_arr C.hseqI' D.hseqI'
C.in_hhom_def C.trg_hcomp D.comp_inv_arr' C.ide_hcomp
by (metis C.ideD(1) D.hcomp_simps(4) cmp_simps(1) cmp_simps(3-5)
\<rho>'.leg0_simps(3) \<rho>'.base_simps(2,4) \<rho>.ide_leg0 \<rho>.ide_base
\<rho>.leg0_simps(3) hseq_fw)
moreover have "(D.inv (\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D \<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w') =
F \<rho> \<star>\<^sub>D F w'"
using w' D.comp_inv_arr' D.comp_cod_arr hseq_fw by auto
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D
(\<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D (D.inv (\<Phi> (r, f \<star>\<^sub>C w'))) \<cdot>\<^sub>D
(\<Phi> (r, f \<star>\<^sub>C w')) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D
(D.inv (\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D
(\<Phi> (r \<star>\<^sub>C f, w')) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w')) \<cdot>\<^sub>D
D.inv (\<Phi> (g, w'))) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
proof -
have "(D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u)) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>') = F r \<star>\<^sub>D F \<theta>'"
using assms(1) \<theta>' D.comp_cod_arr hseq_ru D.comp_inv_arr' by auto
thus ?thesis
using D.comp_assoc by metis
qed
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D
(F (r \<star>\<^sub>C \<theta>') \<cdot>\<^sub>D F \<a>\<^sub>C[r, f, w'] \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w')) \<cdot>\<^sub>D
F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
proof -
have "F (r \<star>\<^sub>C \<theta>') = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>') \<cdot>\<^sub>D D.inv (\<Phi> (r, f \<star>\<^sub>C w'))"
using w' \<theta>' preserves_hcomp hseq_ru by auto
moreover have "F \<a>\<^sub>C[r, f, w'] =
\<Phi> (r, f \<star>\<^sub>C w') \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w'] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w') \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w'))"
using w' preserves_assoc(1) hseq_fw by force
moreover have
"F (\<rho> \<star>\<^sub>C w') = \<Phi> (r \<star>\<^sub>C f, w') \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w') \<cdot>\<^sub>D D.inv (\<Phi> (g, w'))"
using w' preserves_hcomp hseq_fw by fastforce
ultimately show ?thesis
using D.comp_assoc by auto
qed
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F (\<rho>.composite_cell w' \<theta>') \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"
using w' \<theta>' C.comp_assoc hseq_ru hseq_fw by auto
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D (F (\<rho>.composite_cell w' \<theta>') \<cdot>\<^sub>D F \<beta>) \<cdot>\<^sub>D \<Phi> (g, w)"
using D.comp_assoc by simp
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F (\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<beta>) \<cdot>\<^sub>D \<Phi> (g, w)"
proof -
have "F (\<rho>.composite_cell w' \<theta>') \<cdot>\<^sub>D F \<beta> = F (\<rho>.composite_cell w' \<theta>' \<cdot>\<^sub>C \<beta>)"
using w w' \<theta>' \<beta> \<rho>.composite_cell_in_hom
preserves_comp [of "\<rho>.composite_cell w' \<theta>'" \<beta>]
by (metis C.dom_comp C.hcomp_simps(3) C.ide_char C.in_homE C.seqE C.seqI
D.ext D.seqE \<rho>.tab_simps(4) is_extensional preserves_reflects_arr)
thus ?thesis by simp
qed
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D F (\<rho>.composite_cell w \<theta>) \<cdot>\<^sub>D \<Phi> (g, w)"
using eq by simp
also have "... = D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D
F (r \<star>\<^sub>C \<theta>) \<cdot>\<^sub>D F \<a>\<^sub>C[r, f, w] \<cdot>\<^sub>D F (\<rho> \<star>\<^sub>C w) \<cdot>\<^sub>D \<Phi> (g, w)"
using w \<theta> C.comp_assoc hseq_ru hseq_fw D.comp_assoc by auto
also have "... = ((D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D
\<Phi> (r, u)) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>)) \<cdot>\<^sub>D ((D.inv (\<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D
\<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w))) \<cdot>\<^sub>D
\<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D (D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D
((D.inv (\<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D
\<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w)) \<cdot>\<^sub>D D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D \<Phi> (g, w)"
proof -
have "F (r \<star>\<^sub>C \<theta>) = \<Phi> (r, u) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D D.inv (\<Phi> (r, f \<star>\<^sub>C w))"
using w \<theta> preserves_hcomp hseq_ru by auto
moreover have "F \<a>\<^sub>C[r, f, w] =
\<Phi> (r, f \<star>\<^sub>C w) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (r \<star>\<^sub>C f, w))"
using w preserves_assoc(1) hseq_fw by force
moreover have
"F (\<rho> \<star>\<^sub>C w) = \<Phi> (r \<star>\<^sub>C f, w) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
using w preserves_hcomp hseq_fw by fastforce
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (F r \<star>\<^sub>D F \<theta>) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) \<cdot>\<^sub>D \<a>\<^sub>D[F r, F f, F w] \<cdot>\<^sub>D
(D.inv (\<Phi> (r, f)) \<star>\<^sub>D F w) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w)"
proof -
have "(D.inv (\<Phi> (r, u)) \<cdot>\<^sub>D \<Phi> (r, u)) \<cdot>\<^sub>D (F r \<star>\<^sub>D F \<theta>) = F r \<star>\<^sub>D F \<theta>"
using \<theta> D.comp_cod_arr hseq_ru D.comp_inv_arr' by auto
moreover have
"(D.inv (\<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D \<Phi> (r, f \<star>\<^sub>C w)) \<cdot>\<^sub>D (F r \<star>\<^sub>D \<Phi> (f, w)) =
F r \<star>\<^sub>D \<Phi> (f, w)"
using w cmp_components_are_iso D.comp_cod_arr C.hseqI' D.hseqI'
C.in_hhom_def C.trg_hcomp D.comp_inv_arr' C.ide_hcomp
by (metis C.ideD(1) D.hcomp_simps(4) cmp_simps(1) cmp_simps(3-5)
\<rho>'.leg0_simps(3) \<rho>'.base_simps(2,4) \<rho>.ide_leg0 \<rho>.ide_base
\<rho>.leg0_simps(3) hseq_fw)
moreover have "(D.inv (\<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D \<Phi> (r \<star>\<^sub>C f, w)) \<cdot>\<^sub>D (F \<rho> \<star>\<^sub>D F w) =
F \<rho> \<star>\<^sub>D F w"
using w D.comp_inv_arr' D.comp_cod_arr hseq_fw by simp
moreover have "(F \<rho> \<star>\<^sub>D F w) \<cdot>\<^sub>D D.inv (\<Phi> (g, w)) \<cdot>\<^sub>D \<Phi> (g, w) = F \<rho> \<star>\<^sub>D F w"
using w \<theta> D.comp_arr_dom D.comp_inv_arr' hseq_gw by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<rho>'.composite_cell (F w) (F \<theta> \<cdot>\<^sub>D \<Phi> (f, w))"
using w \<theta> 1 D.whisker_left D.whisker_right D.comp_assoc by auto
finally show ?thesis by simp
qed
ultimately show ?thesis
using w w' \<theta> \<theta>' \<beta> eq
\<rho>'.T2 [of "F w" "F w'" "F \<theta> \<cdot>\<^sub>D \<Phi> (f, w)" "F u" "F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')"
"D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w)"]
by blast
qed
obtain \<gamma>' where \<gamma>': "\<guillemotleft>\<gamma>' : F w \<Rightarrow>\<^sub>D F w'\<guillemotright> \<and>
D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) = F g \<star>\<^sub>D \<gamma>' \<and>
F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) = (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>')"
using * by auto
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> F \<gamma> = \<gamma>'"
using \<theta> \<theta> w w' \<gamma>' locally_full [of w w' \<gamma>']
by (metis C.hseqI' C.ideD(1) C.src_hcomp C.vconn_implies_hpar(3)
\<rho>.leg0_simps(2) \<theta>' hseq_fw)
have "\<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)"
proof -
have "F \<theta> = F (\<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>))"
proof -
have "F \<theta> = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D \<gamma>') \<cdot>\<^sub>D D.inv (\<Phi> (f, w))"
using w' \<theta>' \<gamma>' preserves_hcomp hseq_fw D.comp_assoc D.invert_side_of_triangle
by (metis C.in_homE D.comp_arr_inv' cmp_components_are_iso cmp_simps(5)
\<rho>.ide_leg0 \<theta> as_nat_trans.is_natural_1 w)
also have "... = F \<theta>' \<cdot>\<^sub>D F (f \<star>\<^sub>C \<gamma>)"
using w' D.comp_assoc hseq_fw preserves_hcomp cmp_components_are_iso
D.comp_arr_inv'
by (metis C.hseqI' C.in_homE C.trg_cod \<gamma> \<rho>.leg0_in_hom(2))
also have "... = F (\<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>))"
using \<gamma> \<theta> \<theta>' hseq_fw C.hseqI' preserves_comp by force
finally show ?thesis by simp
qed
moreover have "C.par \<theta> (\<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>))"
using \<gamma> \<theta> \<theta>' hseq_fw by fastforce
ultimately show ?thesis
using is_faithful by blast
qed
moreover have "\<beta> = g \<star>\<^sub>C \<gamma>"
proof -
have "F \<beta> = F (g \<star>\<^sub>C \<gamma>)"
proof -
have "F \<beta> = \<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D \<gamma>') \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
by (metis (no_types) C.in_homE D.comp_arr_inv' D.comp_assoc
cmp_components_are_iso cmp_simps(5) \<beta> \<gamma>' \<rho>.ide_leg1 hseq_gw
as_nat_trans.is_natural_1 as_nat_trans.naturality w w')
also have "... = F (g \<star>\<^sub>C \<gamma>)"
using w \<gamma> \<gamma>' preserves_hcomp hseq_gw
by (metis C.hseqE C.hseqI' C.in_homE C.seqE \<open>\<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)\<close>
\<rho>.leg1_simps(2) \<rho>.leg1_simps(5) \<rho>.leg1_simps(6) \<theta> hseq_fw)
finally show ?thesis by simp
qed
moreover have "C.par \<beta> (g \<star>\<^sub>C \<gamma>)"
proof (intro conjI)
show "C.arr \<beta>"
using \<beta> by blast
show 1: "C.hseq g \<gamma>"
using \<gamma> hseq_gw by fastforce
show "C.dom \<beta> = C.dom (g \<star>\<^sub>C \<gamma>)"
using \<gamma> \<beta> 1 by fastforce
show "C.cod \<beta> = C.cod (g \<star>\<^sub>C \<gamma>)"
using \<gamma> \<beta> 1 by fastforce
qed
ultimately show ?thesis
using is_faithful by blast
qed
ultimately have "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma> \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)"
using \<gamma> by blast
moreover have "\<And>\<gamma>\<^sub>1 \<gamma>\<^sub>2. \<guillemotleft>\<gamma>\<^sub>1 : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma>\<^sub>1 \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>1) \<Longrightarrow>
\<guillemotleft>\<gamma>\<^sub>2 : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma>\<^sub>2 \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>2) \<Longrightarrow> \<gamma>\<^sub>1 = \<gamma>\<^sub>2"
proof -
fix \<gamma>\<^sub>1 \<gamma>\<^sub>2
assume \<gamma>\<^sub>1: "\<guillemotleft>\<gamma>\<^sub>1 : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma>\<^sub>1 \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>1)"
assume \<gamma>\<^sub>2: "\<guillemotleft>\<gamma>\<^sub>2 : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma>\<^sub>2 \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>2)"
have F\<beta>\<^sub>1: "F \<beta> = \<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>1) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
using w w' \<beta> hseq_gw \<gamma>\<^sub>1 preserves_hcomp [of g \<gamma>\<^sub>1] cmp_components_are_iso
by auto
have F\<beta>\<^sub>2: "F \<beta> = \<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>2) \<cdot>\<^sub>D D.inv (\<Phi> (g, w))"
using w w' \<beta> hseq_gw \<gamma>\<^sub>2 preserves_hcomp [of g \<gamma>\<^sub>2] cmp_components_are_iso
by auto
have "D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) = F g \<star>\<^sub>D F \<gamma>\<^sub>1"
proof -
have "F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) = \<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>1)"
using w w' \<beta> hseq_gw \<gamma>\<^sub>1 F\<beta>\<^sub>1 preserves_hcomp cmp_components_are_iso
D.invert_side_of_triangle D.iso_inv_iso
by (metis C.arrI D.comp_assoc D.inv_inv \<rho>.ide_leg1 preserves_reflects_arr)
thus ?thesis
using w w' \<beta> hseq_gw \<gamma>\<^sub>1 preserves_hcomp cmp_components_are_iso
D.invert_side_of_triangle
by (metis C.arrI D.cod_comp D.seqE D.seqI F\<beta>\<^sub>1 \<rho>.ide_leg1 preserves_arr)
qed
moreover have "D.inv (\<Phi> (g, w')) \<cdot>\<^sub>D F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) = F g \<star>\<^sub>D F \<gamma>\<^sub>2"
proof -
have "F \<beta> \<cdot>\<^sub>D \<Phi> (g, w) = \<Phi> (g, w') \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<gamma>\<^sub>2)"
using w w' \<beta> hseq_gw \<gamma>\<^sub>2 F\<beta>\<^sub>2 preserves_hcomp cmp_components_are_iso
D.invert_side_of_triangle D.iso_inv_iso
by (metis C.arrI D.comp_assoc D.inv_inv \<rho>.ide_leg1 preserves_reflects_arr)
thus ?thesis
using w w' \<beta> hseq_gw \<gamma>\<^sub>2 preserves_hcomp cmp_components_are_iso
D.invert_side_of_triangle
by (metis C.arrI D.cod_comp D.seqE D.seqI F\<beta>\<^sub>2 \<rho>.ide_leg1 preserves_arr)
qed
moreover have "F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) = (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>1)"
proof -
have "F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) = F (\<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>1)) \<cdot>\<^sub>D \<Phi> (f, w)"
using \<gamma>\<^sub>1 by blast
also have "... = (F \<theta>' \<cdot>\<^sub>D F (f \<star>\<^sub>C \<gamma>\<^sub>1)) \<cdot>\<^sub>D \<Phi> (f, w)"
using \<gamma>\<^sub>1 \<theta> by auto
also have
"... = (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>1) \<cdot>\<^sub>D D.inv (\<Phi> (f, w))) \<cdot>\<^sub>D \<Phi> (f, w)"
using \<gamma>\<^sub>1 hseq_fw preserves_hcomp by auto
also have
"... = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>1) \<cdot>\<^sub>D D.inv (\<Phi> (f, w)) \<cdot>\<^sub>D \<Phi> (f, w)"
using D.comp_assoc by simp
also have "... = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>1) \<cdot>\<^sub>D (F f \<star>\<^sub>D F w)"
by (simp add: D.comp_inv_arr' hseq_fw w)
also have "... = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>1)"
using w \<gamma>\<^sub>1 D.whisker_left [of "F f" "F \<gamma>\<^sub>1" "F w"] D.comp_arr_dom by auto
finally show ?thesis
using D.comp_assoc by simp
qed
moreover have "F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) = (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w')) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>2)"
proof -
have "F \<theta> \<cdot>\<^sub>D \<Phi> (f, w) = F (\<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>\<^sub>2)) \<cdot>\<^sub>D \<Phi> (f, w)"
using \<gamma>\<^sub>2 by blast
also have "... = (F \<theta>' \<cdot>\<^sub>D F (f \<star>\<^sub>C \<gamma>\<^sub>2)) \<cdot>\<^sub>D \<Phi> (f, w)"
using \<gamma>\<^sub>2 \<theta> by auto
also have
"... = (F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>2) \<cdot>\<^sub>D D.inv (\<Phi> (f, w))) \<cdot>\<^sub>D \<Phi> (f, w)"
using \<gamma>\<^sub>2 hseq_fw preserves_hcomp by auto
also have
"... = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>2) \<cdot>\<^sub>D D.inv (\<Phi> (f, w)) \<cdot>\<^sub>D \<Phi> (f, w)"
using D.comp_assoc by simp
also have "... = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>2) \<cdot>\<^sub>D (F f \<star>\<^sub>D F w)"
by (simp add: D.comp_inv_arr' hseq_fw w)
also have "... = F \<theta>' \<cdot>\<^sub>D \<Phi> (f, w') \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<gamma>\<^sub>2)"
using w \<gamma>\<^sub>2 D.whisker_left [of "F f" "F \<gamma>\<^sub>2" "F w"] D.comp_arr_dom by auto
finally show ?thesis
using D.comp_assoc by simp
qed
ultimately have "F \<gamma>\<^sub>1 = F \<gamma>\<^sub>2"
using \<gamma>\<^sub>1 \<gamma>\<^sub>2 * by blast
thus "\<gamma>\<^sub>1 = \<gamma>\<^sub>2"
using \<gamma>\<^sub>1 \<gamma>\<^sub>2 is_faithful [of \<gamma>\<^sub>1 \<gamma>\<^sub>2] by auto
qed
ultimately show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow>\<^sub>C w'\<guillemotright> \<and> \<beta> = g \<star>\<^sub>C \<gamma> \<and> \<theta> = \<theta>' \<cdot>\<^sub>C (f \<star>\<^sub>C \<gamma>)"
by blast
qed
qed
qed
show ?thesis ..
qed
end
end
diff --git a/thys/Category3/HFSetCat.thy b/thys/Category3/HFSetCat.thy
--- a/thys/Category3/HFSetCat.thy
+++ b/thys/Category3/HFSetCat.thy
@@ -1,2259 +1,2261 @@
(* Title: HFSetCat
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2020
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter "The Category of Hereditarily Finite Sets"
theory HFSetCat
imports CategoryWithFiniteLimits CartesianClosedCategory HereditarilyFinite.HF
begin
text\<open>
This theory constructs a category whose objects are in bijective correspondence with
the hereditarily finite sets and whose arrows correspond to the functions between such
sets. We show that this category is cartesian closed and has finite limits.
Note that up to this point we have not constructed any other interpretation for the
@{locale cartesian_closed_category} locale, but it is important to have one to ensure
that the locale assumptions are consistent.
\<close>
section "Preliminaries"
text\<open>
We begin with some preliminary definitions and facts about hereditarily finite sets,
which are better targeted toward what we are trying to do here than what already exists
in @{theory HereditarilyFinite.HF}.
\<close>
text\<open>
The following defines when a hereditarily finite set \<open>F\<close> represents a function from
a hereditarily finite set \<open>B\<close> to a hereditarily finite set \<open>C\<close>. Specifically, \<open>F\<close>
must be a relation from \<open>B\<close> to \<open>C\<close>, whose domain is \<open>B\<close>, whose range is contained in \<open>C\<close>,
and which is single-valued on its domain.
\<close>
definition hfun
where "hfun B C F \<equiv> F \<le> B * C \<and> hfunction F \<and> hdomain F = B \<and> hrange F \<le> C"
lemma hfunI [intro]:
assumes "F \<le> A * B"
and "\<And>X. X \<^bold>\<in> A \<Longrightarrow> \<exists>!Y. \<langle>X, Y\<rangle> \<^bold>\<in> F"
and "\<And>X Y. \<langle>X, Y\<rangle> \<^bold>\<in> F \<Longrightarrow> Y \<^bold>\<in> B"
shows "hfun A B F"
unfolding hfun_def
using assms hfunction_def hrelation_def is_hpair_def hrange_def hconverse_def hdomain_def
apply (intro conjI)
apply auto
by fast
lemma hfunE [elim]:
assumes "hfun B C F"
and "(\<And>Y. Y \<^bold>\<in> B \<Longrightarrow> (\<exists>!Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F) \<and> (\<forall>Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F \<longrightarrow> Z \<^bold>\<in> C)) \<Longrightarrow> T"
shows T
proof -
have "\<And>Y. Y \<^bold>\<in> B \<Longrightarrow> (\<exists>!Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F) \<and> (\<forall>Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F \<longrightarrow> Z \<^bold>\<in> C)"
proof (intro allI impI conjI)
fix Y
assume Y: "Y \<^bold>\<in> B"
show "\<exists>!Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F"
proof -
have "\<exists>Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F"
using assms Y hfun_def hdomain_def by auto
moreover have "\<And>Z Z'. \<lbrakk> \<langle>Y, Z\<rangle> \<^bold>\<in> F; \<langle>Y, Z'\<rangle> \<^bold>\<in> F \<rbrakk> \<Longrightarrow> Z = Z'"
using assms hfun_def hfunction_def by simp
ultimately show ?thesis by blast
qed
show "\<And>Z. \<langle>Y, Z\<rangle> \<^bold>\<in> F \<Longrightarrow> Z \<^bold>\<in> C"
using assms Y hfun_def by auto
qed
thus ?thesis
using assms(2) by simp
qed
text\<open>
The hereditarily finite set \<open>hexp B C\<close> represents the collection of all functions
from \<open>B\<close> to \<open>C\<close>.
\<close>
definition hexp
where "hexp B C = \<lbrace>F \<^bold>\<in> HPow (B * C). hfun B C F\<rbrace>"
lemma hfun_in_hexp:
assumes "hfun B C F"
shows "F \<^bold>\<in> hexp B C"
using assms by (simp add: hexp_def hfun_def)
text\<open>
The function \<open>happ\<close> applies a function \<open>F\<close> from \<open>B\<close> to \<open>C\<close> to an element of \<open>B\<close>,
yielding an element of \<open>C\<close>.
\<close>
abbreviation happ
where "happ \<equiv> app"
lemma happ_mapsto:
assumes "F \<^bold>\<in> hexp B C" and "Y \<^bold>\<in> B"
shows "happ F Y \<^bold>\<in> C" and "happ F Y \<^bold>\<in> hrange F"
proof -
show "happ F Y \<^bold>\<in> C"
using assms app_def hexp_def app_equality hdomain_def hfun_def by auto
show "happ F Y \<^bold>\<in> hrange F"
proof -
have "\<langle>Y, happ F Y\<rangle> \<^bold>\<in> F"
using assms app_def hexp_def app_equality hdomain_def hfun_def by auto
thus ?thesis
using hdomain_def hrange_def hconverse_def by auto
qed
qed
lemma happ_expansion:
assumes "hfun B C F"
shows "F = \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace>"
proof
fix XY
show "XY \<^bold>\<in> F \<longleftrightarrow> XY \<^bold>\<in> \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace>"
proof
show "XY \<^bold>\<in> F \<Longrightarrow> XY \<^bold>\<in> \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace>"
proof -
assume XY: "XY \<^bold>\<in> F"
have "XY \<^bold>\<in> B * C"
using assms XY hfun_def by auto
moreover have "hsnd XY = happ F (hfst XY)"
using assms XY hfunE app_def [of F "hfst XY"] the1_equality [of "\<lambda>y. \<langle>hfst XY, y\<rangle> \<^bold>\<in> F"]
calculation
by auto
ultimately show "XY \<^bold>\<in> \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace>" by simp
qed
show "XY \<^bold>\<in> \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace> \<Longrightarrow> XY \<^bold>\<in> F"
proof -
assume XY: "XY \<^bold>\<in> \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace>"
show "XY \<^bold>\<in> F"
using assms XY app_def [of F "hfst XY"] the1_equality [of "\<lambda>y. \<langle>hfst XY, y\<rangle> \<^bold>\<in> F"]
by fastforce
qed
qed
qed
text\<open>
Function \<open>hlam\<close> takes a function \<open>F\<close> from \<open>A * B\<close> to \<open>C\<close> to a function \<open>hlam F\<close>
from \<open>A\<close> to \<open>hexp B C\<close>.
\<close>
definition hlam
where "hlam A B C F =
\<lbrace>XG \<^bold>\<in> A * hexp B C.
\<forall>YZ. YZ \<^bold>\<in> hsnd XG \<longleftrightarrow> is_hpair YZ \<and> \<langle>\<langle>hfst XG, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
lemma hfun_hlam:
assumes "hfun (A * B) C F"
shows "hfun A (hexp B C) (hlam A B C F)"
proof
show "hlam A B C F \<le> A * hexp B C"
using assms hlam_def by auto
show "\<And>X. X \<^bold>\<in> A \<Longrightarrow> \<exists>!Y. \<langle>X, Y\<rangle> \<^bold>\<in> hlam A B C F"
proof
fix X
assume X: "X \<^bold>\<in> A"
let ?G = "\<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
have 1: "?G \<^bold>\<in> hexp B C"
using assms X hexp_def by fastforce
show "\<langle>X, ?G\<rangle> \<^bold>\<in> hlam A B C F"
using assms X 1 is_hpair_def hfun_def hlam_def by auto
fix Y
assume XY: "\<langle>X, Y\<rangle> \<^bold>\<in> hlam A B C F"
show "Y = ?G"
using assms X XY hlam_def hexp_def by fastforce
qed
show "\<And>X Y. \<langle>X, Y\<rangle> \<^bold>\<in> hlam A B C F \<Longrightarrow> Y \<^bold>\<in> hexp B C"
using assms hlam_def hexp_def by simp
qed
lemma happ_hlam:
assumes "X \<^bold>\<in> A" and "hfun (A * B) C F"
shows "\<exists>!G. \<langle>X, G\<rangle> \<^bold>\<in> hlam A B C F"
and "happ (hlam A B C F) X = (THE G. \<langle>X, G\<rangle> \<^bold>\<in> hlam A B C F)"
and "happ (hlam A B C F) X = \<lbrace>yz \<^bold>\<in> B * C. \<langle>\<langle>X, hfst yz\<rangle>, hsnd yz\<rangle> \<^bold>\<in> F\<rbrace>"
and "Y \<^bold>\<in> B \<Longrightarrow> happ (happ (hlam A B C F) X) Y = happ F \<langle>X, Y\<rangle>"
proof -
show 1: "\<exists>!G. \<langle>X, G\<rangle> \<^bold>\<in> hlam A B C F"
using assms(1,2) hfun_hlam hfunE
by (metis (full_types))
show 2: "happ (hlam A B C F) X = (THE G. \<langle>X, G\<rangle> \<^bold>\<in> hlam A B C F)"
using assms app_def by simp
show "happ (happ (hlam A B C F) X) Y = happ F \<langle>X, Y\<rangle>"
proof -
have 3: "\<langle>X, happ (hlam A B C F) X\<rangle> \<^bold>\<in> hlam A B C F"
using assms(1) 1 2 theI' [of "\<lambda>G. \<langle>X, G\<rangle> \<^bold>\<in> hlam A B C F"] by simp
hence "\<exists>!Z. happ (happ (hlam A B C F) X) = Z"
by simp
moreover have "happ (happ (hlam A B C F) X) Y = happ F \<langle>X, Y\<rangle>"
using assms(1-2) 3 hlam_def is_hpair_def app_def by simp
ultimately show ?thesis by simp
qed
show "happ (hlam A B C F) X = \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
proof -
let ?G = "\<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
have 4: "hfun B C ?G"
proof
show "\<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace> \<le> B * C"
using assms by auto
show "\<And>Y. Y \<^bold>\<in> B \<Longrightarrow> \<exists>!Z. \<langle>Y, Z\<rangle> \<^bold>\<in> \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
proof -
fix Y
assume Y: "Y \<^bold>\<in> B"
have XY: "\<langle>X, Y\<rangle> \<^bold>\<in> A * B"
using assms Y by simp
hence 1: "\<exists>!Z. \<langle>\<langle>X, Y\<rangle>, Z\<rangle> \<^bold>\<in> F"
using assms XY hfunE [of "A * B" C F] by metis
obtain Z where Z: "\<langle>\<langle>X, Y\<rangle>, Z\<rangle> \<^bold>\<in> F"
using 1 by auto
have "\<exists>Z. \<langle>Y, Z\<rangle> \<^bold>\<in> \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
proof -
have "\<langle>Y, Z\<rangle> \<^bold>\<in> B * C"
using assms Y Z by blast
moreover have "\<langle>\<langle>X, hfst \<langle>Y, Z\<rangle>\<rangle>, hsnd \<langle>Y, Z\<rangle>\<rangle> \<^bold>\<in> F"
using assms Y Z by simp
ultimately show ?thesis by auto
qed
moreover have "\<And>Z Z'. \<lbrakk>\<langle>Y, Z\<rangle> \<^bold>\<in> \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>;
\<langle>Y, Z'\<rangle> \<^bold>\<in> \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>\<rbrakk> \<Longrightarrow> Z = Z'"
using assms Y by auto
ultimately show "\<exists>!Z. \<langle>Y, Z\<rangle> \<^bold>\<in> \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace>"
by auto
qed
show "\<And>Y Z. \<langle>Y, Z\<rangle> \<^bold>\<in> \<lbrace>YZ \<^bold>\<in> B * C. \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F\<rbrace> \<Longrightarrow> Z \<^bold>\<in> C"
using assms by simp
qed
have "\<langle>X, ?G\<rangle> \<^bold>\<in> hlam A B C F"
proof -
have "\<langle>X, ?G\<rangle> \<^bold>\<in> A * hexp B C"
using assms 4
by (simp add: hfun_in_hexp)
moreover have "\<forall>YZ. YZ \<^bold>\<in> ?G \<longleftrightarrow> is_hpair YZ \<and> \<langle>\<langle>X, hfst YZ\<rangle>, hsnd YZ\<rangle> \<^bold>\<in> F"
using assms 1 is_hpair_def hfun_def by auto
ultimately show ?thesis
using assms 1 hlam_def by simp
qed
thus "happ (hlam A B C F) X = ?G"
using assms 2 4 app_equality hfun_def hfun_hlam by auto
qed
qed
section "Construction of the Category"
locale hfsetcat
begin
text\<open>
We construct the category of hereditarily finite sets and functions simply by applying
the generic ``set category'' construction, using the hereditarily finite sets as the
universe, and constraining the collections of such sets that determine objects of the
category to those that are finite.
\<close>
interpretation setcat \<open>undefined :: hf\<close> finite
using finite_subset
by unfold_locales blast+
+ interpretation set_category comp \<open>\<lambda>A. A \<subseteq> Collect terminal \<and> finite (elem_of ` A)\<close>
+ using is_set_category by blast
lemma set_ide_char:
shows "A \<in> set ` Collect ide \<longleftrightarrow> A \<subseteq> Univ \<and> finite A"
proof
assume A: "A \<in> set ` Collect ide"
show "A \<subseteq> Univ \<and> finite A"
proof
show "A \<subseteq> Univ"
using A setp_set' by auto
obtain a where a: "ide a \<and> A = set a"
using A by blast
have "finite (elem_of ` set a)"
using a setp_set_ide by blast
moreover have "inj_on elem_of (set a)"
proof -
have "inj_on elem_of Univ"
using bij_elem_of bij_betw_imp_inj_on by auto
moreover have "set a \<subseteq> Univ"
using a setp_set' [of a] by blast
ultimately show ?thesis
using inj_on_subset by auto
qed
ultimately show "finite A"
using a A finite_imageD [of elem_of "set a"] by blast
qed
next
assume A: "A \<subseteq> Univ \<and> finite A"
have "ide (mkIde A)"
using A ide_mkIde by simp
moreover have "set (mkIde A) = A"
using A finite_imp_setp set_mkIde by presburger
ultimately show "A \<in> set ` Collect ide" by blast
qed
lemma set_ideD:
assumes "ide a"
shows "set a \<subseteq> Univ" and "finite (set a)"
using assms set_ide_char by auto
lemma ide_mkIdeI [intro]:
assumes "A \<subseteq> Univ" and "finite A"
shows "ide (mkIde A)" and "set (mkIde A) = A"
using assms ide_mkIde set_mkIde by auto
interpretation category_with_terminal_object comp
using terminal_unity by unfold_locales auto
text\<open>
We verify that the objects of HF are indeed in bijective correspondence with the
hereditarily finite sets.
\<close>
definition ide_to_hf
where "ide_to_hf a = HF (elem_of ` set a)"
definition hf_to_ide
where "hf_to_ide x = mkIde (arr_of ` hfset x)"
lemma ide_to_hf_mapsto:
shows "ide_to_hf \<in> Collect ide \<rightarrow> UNIV"
by simp
lemma hf_to_ide_mapsto:
shows "hf_to_ide \<in> UNIV \<rightarrow> Collect ide"
proof
fix x :: hf
have "finite (arr_of ` hfset x)"
by simp
moreover have "arr_of ` hfset x \<subseteq> Univ"
by (metis (mono_tags, lifting) UNIV_I bij_arr_of bij_betw_def imageE image_eqI subsetI)
ultimately have "ide (mkIde (arr_of ` hfset x))"
using finite_imp_setp ide_mkIde by presburger
thus "hf_to_ide x \<in> Collect ide"
using hf_to_ide_def by simp
qed
lemma hf_to_ide_ide_to_hf:
assumes "a \<in> Collect ide"
shows "hf_to_ide (ide_to_hf a) = a"
proof -
have "hf_to_ide (ide_to_hf a) = mkIde (arr_of ` hfset (HF (elem_of ` set a)))"
using hf_to_ide_def ide_to_hf_def by simp
also have "... = a"
proof -
have "mkIde (arr_of ` hfset (HF (elem_of ` set a))) = mkIde (arr_of ` elem_of ` set a)"
proof -
have "finite (set a)"
using assms set_ide_char by blast
hence "finite (elem_of ` set a)"
by simp
hence "hfset (HF (elem_of ` set a)) = elem_of ` set a"
using hfset_HF [of "elem_of ` set a"] by simp
thus ?thesis by simp
qed
also have "... = a"
proof -
have "set a \<subseteq> Univ"
using assms set_ide_char by blast
hence "\<And>x. x \<in> set a \<Longrightarrow> arr_of (elem_of x) = x"
using assms by auto
hence "arr_of ` elem_of ` set a = set a"
by force
thus ?thesis
using assms ide_char mkIde_set by simp
qed
finally show ?thesis by blast
qed
finally show "hf_to_ide (ide_to_hf a) = a" by blast
qed
lemma ide_to_hf_hf_to_ide:
assumes "x \<in> UNIV"
shows "ide_to_hf (hf_to_ide x) = x"
proof -
have "HF (elem_of ` set (mkIde (arr_of ` hfset x))) = x"
proof -
have "HF (elem_of ` set (mkIde (arr_of ` hfset x))) = HF (elem_of ` arr_of ` hfset x)"
using assms set_mkIde [of "arr_of ` hfset x"] arr_of_mapsto mkIde_def by auto
also have "... = HF (hfset x)"
proof -
have "\<And>A. elem_of ` arr_of ` A = A"
using elem_of_arr_of by force
thus ?thesis by metis
qed
also have "... = x" by simp
finally show ?thesis by blast
qed
thus ?thesis
using assms ide_to_hf_def hf_to_ide_def by simp
qed
lemma bij_betw_ide_hf_set:
shows "bij_betw ide_to_hf (Collect ide) (UNIV :: hf set)"
using ide_to_hf_mapsto hf_to_ide_mapsto ide_to_hf_hf_to_ide hf_to_ide_ide_to_hf
by (intro bij_betwI) auto
lemma ide_implies_finite_set:
assumes "ide a"
shows "finite (set a)" and "finite (hom unity a)"
proof -
show 1: "finite (set a)"
using assms set_ide_char by blast
show "finite (hom unity a)"
using assms 1 bij_betw_points_and_set finite_imageD inj_img set_def by auto
qed
text\<open>
We establish the connection between the membership relation defined for hereditarily
finite sets and the corresponding membership relation associated with the set category.
\<close>
lemma arr_of_membI [intro]:
assumes "x \<^bold>\<in> ide_to_hf a"
shows "arr_of x \<in> set a"
proof -
let ?X = "inv_into (set a) elem_of x"
have "x = elem_of ?X \<and> ?X \<in> set a"
using assms
by (simp add: f_inv_into_f ide_to_hf_def inv_into_into)
thus ?thesis
by (metis (no_types, lifting) arr_of_elem_of elem_set_implies_incl_in
elem_set_implies_set_eq_singleton incl_in_def mem_Collect_eq terminal_char2)
qed
lemma elem_of_membI [intro]:
assumes "ide a" and "x \<in> set a"
shows "elem_of x \<^bold>\<in> ide_to_hf a"
proof -
have "finite (elem_of ` set a)"
using assms ide_implies_finite_set [of a] by simp
hence "elem_of x \<in> hfset (ide_to_hf a)"
using assms ide_to_hf_def hfset_HF [of "elem_of ` set a"] by simp
thus ?thesis
using hmem_def by blast
qed
text\<open>
We show that each hom-set \<open>hom a b\<close> is in bijective correspondence with
the elements of the hereditarily finite set \<open>hfun (ide_to_hf a) (ide_to_hf b)\<close>.
\<close>
definition arr_to_hfun
where "arr_to_hfun f = \<lbrace>XY \<^bold>\<in> ide_to_hf (dom f) * ide_to_hf (cod f).
hsnd XY = elem_of (Fun f (arr_of (hfst XY)))\<rbrace>"
definition hfun_to_arr
where "hfun_to_arr B C F =
mkArr (arr_of ` hfset B) (arr_of ` hfset C) (\<lambda>x. arr_of (happ F (elem_of x)))"
lemma hfun_arr_to_hfun:
assumes "arr f"
shows "hfun (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f)"
proof
show "arr_to_hfun f \<le> ide_to_hf (dom f) * ide_to_hf (cod f)"
using assms arr_to_hfun_def by auto
show "\<And>X. X \<^bold>\<in> ide_to_hf (dom f) \<Longrightarrow> \<exists>!Y. \<langle>X, Y\<rangle> \<^bold>\<in> arr_to_hfun f"
proof
fix X
assume X: "X \<^bold>\<in> ide_to_hf (dom f)"
show "\<langle>X, elem_of (Fun f (arr_of X))\<rangle> \<^bold>\<in> arr_to_hfun f"
proof -
have "\<langle>X, elem_of (Fun f (arr_of X))\<rangle> \<^bold>\<in> \<lbrace>XY \<^bold>\<in> ide_to_hf (dom f) * ide_to_hf (cod f).
hsnd XY = elem_of (Fun f (arr_of (hfst XY)))\<rbrace>"
proof -
have "hsnd \<langle>X, elem_of (Fun f (arr_of X))\<rangle> =
elem_of (Fun f (arr_of (hfst \<langle>X, elem_of (Fun f (arr_of X))\<rangle>)))"
using assms X by simp
moreover have "\<langle>X, elem_of (Fun f (arr_of X))\<rangle> \<^bold>\<in> ide_to_hf (dom f) * ide_to_hf (cod f)"
proof -
have "elem_of (Fun f (arr_of X)) \<^bold>\<in> ide_to_hf (cod f)"
proof (intro elem_of_membI)
show "ide (cod f)"
using assms ide_cod by simp
show "Fun f (arr_of X) \<in> Cod f"
using assms X Fun_mapsto arr_of_membI by auto
qed
thus ?thesis
using X by simp
qed
ultimately show ?thesis by simp
qed
thus ?thesis
using arr_to_hfun_def by simp
qed
fix Y
assume XY: "\<langle>X, Y\<rangle> \<^bold>\<in> arr_to_hfun f"
show "Y = elem_of (Fun f (arr_of X))"
using assms X XY arr_to_hfun_def by auto
qed
show "\<And>X Y. \<langle>X, Y\<rangle> \<^bold>\<in> arr_to_hfun f \<Longrightarrow> Y \<^bold>\<in> ide_to_hf (cod f)"
using assms arr_to_hfun_def ide_to_hf_def
\<open>arr_to_hfun f \<le> ide_to_hf (dom f) * ide_to_hf (cod f)\<close>
by blast
qed
lemma arr_to_hfun_in_hexp:
assumes "arr f"
shows "arr_to_hfun f \<^bold>\<in> hexp (ide_to_hf (dom f)) (ide_to_hf (cod f))"
using assms arr_to_hfun_def hfun_arr_to_hfun hexp_def by auto
lemma hfun_to_arr_in_hom:
assumes "hfun B C F"
shows "\<guillemotleft>hfun_to_arr B C F : hf_to_ide B \<rightarrow> hf_to_ide C\<guillemotright>"
proof
let ?f = "mkArr (arr_of ` hfset B) (arr_of ` hfset C) (\<lambda>x. arr_of (happ F (elem_of x)))"
have 0: "arr ?f"
proof -
have "arr_of ` hfset B \<subseteq> Univ \<and> arr_of ` hfset C \<subseteq> Univ"
using arr_of_mapsto by auto
moreover have "(\<lambda>x. arr_of (happ F (elem_of x))) \<in> arr_of ` hfset B \<rightarrow> arr_of ` hfset C"
proof
fix x
assume x: "x \<in> arr_of ` hfset B"
have "happ F (elem_of x) \<in> hfset C"
using assms x happ_mapsto hfun_in_hexp
by (metis elem_of_arr_of HF_hfset finite_hfset hmem_HF_iff imageE)
thus "arr_of (happ F (elem_of x)) \<in> arr_of ` hfset C"
by simp
qed
ultimately show ?thesis
using arr_mkArr
by (meson finite_hfset finite_iff_ordLess_natLeq finite_imageI)
qed
show 1: "arr (hfun_to_arr B C F)"
using 0 hfun_to_arr_def by simp
show "dom (hfun_to_arr B C F) = hf_to_ide B"
using 1 hfun_to_arr_def hf_to_ide_def dom_mkArr by auto
show "cod (hfun_to_arr B C F) = hf_to_ide C"
using 1 hfun_to_arr_def hf_to_ide_def cod_mkArr by auto
qed
text\<open>
The comprehension notation from @{theory HereditarilyFinite.HF} interferes in an
unfortunate way with the restriction notation from @{theory "HOL-Library.FuncSet"},
making it impossible to use both in the present context.
\<close>
lemma Fun_char:
assumes "arr f"
shows "Fun f = restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f)"
proof
fix x
show "Fun f x = restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) x"
proof (cases "x \<in> Dom f")
show "x \<notin> Dom f \<Longrightarrow> ?thesis"
using assms Fun_mapsto Fun_def restrict_apply by simp
show "x \<in> Dom f \<Longrightarrow> ?thesis"
proof -
assume x: "x \<in> Dom f"
have 1: "hfun (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f)"
using assms app_def arr_to_hfun_def hfun_arr_to_hfun
the1_equality [of "\<lambda>y. \<langle>elem_of x, y\<rangle> \<^bold>\<in> arr_to_hfun f" "elem_of (Fun f x)"]
by simp
have 2: "\<exists>!Y. \<langle>elem_of x, Y\<rangle> \<^bold>\<in> arr_to_hfun f"
using assms x 1 hfunE elem_of_membI ide_dom
by (metis (no_types, lifting))
have "Fun f x = arr_of (elem_of (Fun f x))"
proof -
have "Fun f x \<in> Univ"
using assms x ide_cod Fun_mapsto [of f] set_ide_char by blast
thus ?thesis
using arr_of_elem_of by simp
qed
also have "... = arr_of (happ (arr_to_hfun f) (elem_of x))"
proof -
have "\<langle>elem_of x, elem_of (Fun f x)\<rangle> \<^bold>\<in> arr_to_hfun f"
proof -
have "\<langle>elem_of x, elem_of (Fun f x)\<rangle> \<^bold>\<in> ide_to_hf (dom f) * ide_to_hf (cod f)"
using assms x ide_dom ide_cod Fun_mapsto by fast
moreover have "elem_of (Fun f x) = elem_of (Fun f (arr_of (elem_of x)))"
by (metis (no_types, lifting) arr_of_elem_of setp_set_ide assms ide_dom subsetD x)
ultimately show ?thesis
using arr_to_hfun_def by auto
qed
moreover have "\<langle>elem_of x, happ (arr_to_hfun f) (elem_of x)\<rangle> \<^bold>\<in> arr_to_hfun f"
using assms x 1 2 app_equality hfun_def by blast
ultimately show ?thesis
using 2 by fastforce
qed
also have "... = restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) x"
using assms x ide_dom by auto
finally show ?thesis by simp
qed
qed
qed
lemma Fun_hfun_to_arr:
assumes "hfun B C F"
shows "Fun (hfun_to_arr B C F) = restrict (\<lambda>x. arr_of (happ F (elem_of x))) (arr_of ` hfset B)"
proof -
have "arr (hfun_to_arr B C F)"
using assms hfun_to_arr_in_hom by blast
hence "arr (mkArr (arr_of ` hfset B) (arr_of ` hfset C) (\<lambda>x. arr_of (happ F (elem_of x))))"
using hfun_to_arr_def by simp
thus ?thesis
using assms hfun_to_arr_def Fun_mkArr by simp
qed
lemma arr_of_img_hfset_ide_to_hf:
assumes "ide a"
shows "arr_of ` hfset (ide_to_hf a) = set a"
proof -
have "arr_of ` hfset (ide_to_hf a) = arr_of ` hfset (HF (elem_of ` set a))"
using ide_to_hf_def by simp
also have "... = arr_of ` elem_of ` set a"
using assms ide_implies_finite_set(1) ide_char by auto
also have "... = set a"
proof -
have "\<And>x. x \<in> set a \<Longrightarrow> arr_of (elem_of x) = x"
using assms ide_char arr_of_elem_of setp_set_ide by blast
thus ?thesis by force
qed
finally show ?thesis by blast
qed
lemma hfun_to_arr_arr_to_hfun:
assumes "arr f"
shows "hfun_to_arr (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f) = f"
proof -
have 0: "hfun_to_arr (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f) =
mkArr (arr_of ` hfset (ide_to_hf (dom f))) (arr_of ` hfset (ide_to_hf (cod f)))
(\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x)))"
unfolding hfun_to_arr_def by blast
also have "... = mkArr (Dom f) (Cod f)
(restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f))"
proof (intro mkArr_eqI)
show 1: "arr_of ` hfset (ide_to_hf (dom f)) = Dom f"
using assms arr_of_img_hfset_ide_to_hf ide_dom by simp
show 2: "arr_of ` hfset (ide_to_hf (cod f)) = Cod f"
using assms arr_of_img_hfset_ide_to_hf ide_cod by simp
show "arr (mkArr (arr_of ` hfset (ide_to_hf (dom f))) (arr_of ` hfset (ide_to_hf (cod f)))
(\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))))"
using 0 1 2
by (metis (no_types, lifting) arrI assms hfun_arr_to_hfun hfun_to_arr_in_hom)
show "\<And>x. x \<in> arr_of ` hfset (ide_to_hf (dom f)) \<Longrightarrow>
arr_of (happ (arr_to_hfun f) (elem_of x)) =
restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) x"
using assms 1 by simp
qed
also have "... = mkArr (Dom f) (Cod f) (Fun f)"
using assms Fun_char mkArr_eqI by simp
also have "... = f"
using assms mkArr_Fun by blast
finally show ?thesis by simp
qed
lemma arr_to_hfun_hfun_to_arr:
assumes "hfun B C F"
shows "arr_to_hfun (hfun_to_arr B C F) = F"
proof -
have "arr_to_hfun (hfun_to_arr B C F) =
\<lbrace>XY \<^bold>\<in> ide_to_hf (dom (hfun_to_arr B C F)) * ide_to_hf (cod (hfun_to_arr B C F)).
hsnd XY = elem_of (Fun (hfun_to_arr B C F) (arr_of (hfst XY)))\<rbrace>"
unfolding arr_to_hfun_def by blast
also have
"... = \<lbrace>XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)).
hsnd XY = elem_of (Fun (hfun_to_arr B C F) (arr_of (hfst XY)))\<rbrace>"
using assms hfun_to_arr_in_hom [of B C F] hf_to_ide_def
by (metis (no_types, lifting) in_homE)
also have
"... = \<lbrace>XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)).
hsnd XY = elem_of (restrict (\<lambda>x. arr_of (happ F (elem_of x))) (arr_of ` hfset B)
(arr_of (hfst XY)))\<rbrace>"
using assms Fun_hfun_to_arr by simp
also have
"... = \<lbrace>XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)).
hsnd XY = elem_of (arr_of (happ F (elem_of (arr_of (hfst XY)))))\<rbrace>"
proof -
have
1: "\<And>XY. XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C))
\<Longrightarrow> arr_of (hfst XY) \<in> arr_of ` hfset B"
proof -
fix XY
assume
XY: "XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C))"
have "hfst XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B))"
using XY by auto
thus "arr_of (hfst XY) \<in> arr_of ` hfset B"
using assms arr_of_membI [of "hfst XY" "mkIde (arr_of ` hfset B)"] set_mkIde
by (metis (mono_tags, lifting) arrI arr_mkArr hfun_to_arr_def hfun_to_arr_in_hom)
qed
show ?thesis
proof -
have
"\<And>XY. (XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)) \<and>
hsnd XY = elem_of (restrict (\<lambda>x. arr_of (happ F (elem_of x))) (arr_of ` hfset B)
(arr_of (hfst XY))))
\<longleftrightarrow>
(XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)) \<and>
hsnd XY = elem_of (arr_of (happ F (elem_of (arr_of (hfst XY))))))"
using 1 by auto
thus ?thesis by blast
qed
qed
also have
"... = \<lbrace>XY \<^bold>\<in> ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)).
hsnd XY = happ F (hfst XY)\<rbrace>"
by simp
also have "... = \<lbrace>XY \<^bold>\<in> B * C. hsnd XY = happ F (hfst XY)\<rbrace>"
using assms hf_to_ide_def ide_to_hf_hf_to_ide by force
also have "... = F"
using assms happ_expansion by simp
finally show ?thesis by simp
qed
lemma bij_betw_hom_hfun:
assumes "ide a" and "ide b"
shows "bij_betw arr_to_hfun (hom a b) {F. hfun (ide_to_hf a) (ide_to_hf b) F}"
proof (intro bij_betwI)
show "arr_to_hfun \<in> hom a b \<rightarrow> {F. hfun (ide_to_hf a) (ide_to_hf b) F}"
using assms arr_to_hfun_in_hexp hexp_def hfun_arr_to_hfun by blast
show "hfun_to_arr (ide_to_hf a) (ide_to_hf b)
\<in> {F. hfun (ide_to_hf a) (ide_to_hf b) F} \<rightarrow> hom a b"
using assms hfun_to_arr_in_hom
by (metis (no_types, lifting) Pi_I hf_to_ide_ide_to_hf mem_Collect_eq)
show "\<And>x. x \<in> hom a b \<Longrightarrow> hfun_to_arr (ide_to_hf a) (ide_to_hf b) (arr_to_hfun x) = x"
using assms hfun_to_arr_arr_to_hfun by blast
show "\<And>y. y \<in> {F. hfun (ide_to_hf a) (ide_to_hf b) F} \<Longrightarrow>
arr_to_hfun (hfun_to_arr (ide_to_hf a) (ide_to_hf b) y) = y"
using assms arr_to_hfun_hfun_to_arr by simp
qed
text\<open>
We next relate composition of arrows in the category to the corresponding operation
on hereditarily finite sets.
\<close>
definition hcomp
where "hcomp G F =
\<lbrace>XZ \<^bold>\<in> hdomain F * hrange G. hsnd XZ = happ G (happ F (hfst XZ))\<rbrace>"
lemma hfun_hcomp:
assumes "hfun A B F" and "hfun B C G"
shows "hfun A C (hcomp G F)"
proof
show "hcomp G F \<le> A * C"
using assms hcomp_def hfun_def by auto
show "\<And>X. X \<^bold>\<in> A \<Longrightarrow> \<exists>!Y. \<langle>X, Y\<rangle> \<^bold>\<in> hcomp G F"
proof
fix X
assume X: "X \<^bold>\<in> A"
show "\<langle>X, happ G (happ F X)\<rangle> \<^bold>\<in> hcomp G F"
unfolding hcomp_def
using assms X hfunE happ_mapsto hfun_in_hexp
by (metis (mono_tags, lifting) HCollect_iff hfst_conv hfun_def hsnd_conv timesI)
show "\<And>X Y. \<lbrakk>X \<^bold>\<in> A; \<langle>X, Y\<rangle> \<^bold>\<in> hcomp G F\<rbrakk> \<Longrightarrow> Y = happ G (happ F X)"
unfolding hcomp_def by simp
qed
show "\<And>X Y. \<langle>X, Y\<rangle> \<^bold>\<in> hcomp G F \<Longrightarrow> Y \<^bold>\<in> C"
unfolding hcomp_def
using assms hfunE happ_mapsto hfun_in_hexp
by (metis HCollectE hfun_def hsubsetCE timesD2)
qed
lemma arr_to_hfun_comp:
assumes "seq g f"
shows "arr_to_hfun (comp g f) = hcomp (arr_to_hfun g) (arr_to_hfun f)"
proof -
have 1: "hdomain (arr_to_hfun f) = ide_to_hf (dom f)"
using assms hfun_arr_to_hfun hfun_def by blast
have "arr_to_hfun (comp g f) =
\<lbrace>XZ \<^bold>\<in> ide_to_hf (dom f) * ide_to_hf (cod g).
hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))\<rbrace>"
unfolding arr_to_hfun_def comp_def
using assms by fastforce
also have "... = \<lbrace>XZ \<^bold>\<in> hdomain (arr_to_hfun f) * hrange (arr_to_hfun g).
hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))\<rbrace>"
proof
fix XZ
have "hfst XZ \<^bold>\<in> hdomain (arr_to_hfun f)
\<Longrightarrow> hsnd XZ \<^bold>\<in> ide_to_hf (cod g) \<and>
hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))
\<longleftrightarrow>
hsnd XZ \<^bold>\<in> hrange (arr_to_hfun g) \<and>
hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))"
proof
assume XZ: "hfst XZ \<^bold>\<in> hdomain (arr_to_hfun f)"
have 2: "arr_of (hfst XZ) \<in> Dom f"
using XZ 1 hfsetcat.arr_of_membI by auto
have 3: "arr_of (happ (arr_to_hfun f) (hfst XZ)) \<in> Dom g"
using assms XZ 2
by (metis (no_types, lifting) "1" happ_mapsto(1) hfsetcat.arr_of_membI
arr_to_hfun_in_hexp seqE)
have 4: "elem_of (Fun (comp g f) (arr_of (hfst XZ))) =
happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))"
proof -
have "elem_of (Fun (comp g f) (arr_of (hfst XZ))) =
elem_of (restrict (Fun g o Fun f) (Dom f) (arr_of (hfst XZ)))"
using assms Fun_comp Fun_char by simp
also have "... = elem_of ((Fun g o Fun f) (arr_of (hfst XZ)))"
using XZ 2 by auto
also have "... = elem_of (Fun g (Fun f (arr_of (hfst XZ))))"
by simp
also have
"... = elem_of (Fun g (restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f)
(arr_of (hfst XZ))))"
proof -
have "Fun f = restrict (\<lambda>x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f)"
using assms Fun_char [of f] by blast
thus ?thesis by simp
qed
also have "... = elem_of (Fun g (arr_of (happ (arr_to_hfun f) (hfst XZ))))"
using 2 by simp
also have "... = elem_of (restrict (\<lambda>x. arr_of (happ (arr_to_hfun g) (elem_of x))) (Dom g)
(arr_of (happ (arr_to_hfun f) (hfst XZ))))"
proof -
have "Fun g = restrict (\<lambda>x. arr_of (happ (arr_to_hfun g) (elem_of x))) (Dom g)"
using assms Fun_char [of g] by blast
thus ?thesis by simp
qed
also have "... = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))"
using 3 by simp
finally show ?thesis by blast
qed
have 5: "elem_of (Fun (comp g f) (arr_of (hfst XZ))) \<^bold>\<in> hrange (arr_to_hfun g)"
proof -
have "happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ)) \<^bold>\<in> hrange (arr_to_hfun g)"
using assms 1 3 XZ hfun_arr_to_hfun happ_mapsto arr_to_hfun_in_hexp arr_to_hfun_def
by (metis (no_types, lifting) seqE)
thus ?thesis
using XZ 4 by simp
qed
show "hsnd XZ \<^bold>\<in> ide_to_hf (cod g) \<and>
hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))
\<Longrightarrow>
hsnd XZ \<^bold>\<in> hrange (arr_to_hfun g) \<and>
hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))"
using XZ 4 5 by simp
show "hsnd XZ \<^bold>\<in> hrange (arr_to_hfun g) \<and>
hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))
\<Longrightarrow>
hsnd XZ \<^bold>\<in> ide_to_hf (cod g) \<and>
hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))"
using assms XZ 1 4
by (metis (no_types, lifting) arr_to_hfun_in_hexp happ_mapsto(1) seqE)
qed
thus "XZ \<^bold>\<in> \<lbrace>XZ \<^bold>\<in> ide_to_hf (dom f) * ide_to_hf (cod g).
hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))\<rbrace>
\<longleftrightarrow>
XZ \<^bold>\<in> \<lbrace>XZ \<^bold>\<in> hdomain (arr_to_hfun f) * hrange (arr_to_hfun g).
hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))\<rbrace>"
using 1 is_hpair_def by auto
qed
also have "... = hcomp (arr_to_hfun g) (arr_to_hfun f)"
using assms arr_to_hfun_def hcomp_def by simp
finally show ?thesis by simp
qed
lemma hfun_to_arr_hcomp:
assumes "hfun A B F" and "hfun B C G"
shows "hfun_to_arr A C (hcomp G F) = comp (hfun_to_arr B C G) (hfun_to_arr A B F)"
proof -
have 1: "arr_to_hfun (hfun_to_arr A C (hcomp G F)) =
arr_to_hfun (comp (hfun_to_arr B C G) (hfun_to_arr A B F))"
proof -
have "arr_to_hfun (comp (hfun_to_arr B C G) (hfun_to_arr A B F)) =
hcomp (arr_to_hfun (hfun_to_arr B C G)) (arr_to_hfun (hfun_to_arr A B F))"
using assms arr_to_hfun_comp hfun_to_arr_in_hom by blast
also have "... = hcomp G F"
using assms by (simp add: arr_to_hfun_hfun_to_arr)
also have "... = arr_to_hfun (hfun_to_arr A C (hcomp G F))"
proof -
have "hfun A C (hcomp G F)"
using assms hfun_hcomp by simp
thus ?thesis
by (simp add: arr_to_hfun_hfun_to_arr)
qed
finally show ?thesis by simp
qed
show ?thesis
proof -
have "hfun_to_arr A C (hcomp G F) \<in> hom (hf_to_ide A) (hf_to_ide C)"
using assms hfun_hcomp hf_to_ide_def hfun_to_arr_in_hom by auto
moreover have "comp (hfun_to_arr B C G) (hfun_to_arr A B F)
\<in> hom (hf_to_ide A) (hf_to_ide C)"
using assms hfun_to_arr_in_hom hf_to_ide_def
by (metis (no_types, lifting) comp_in_homI mem_Collect_eq)
moreover have "inj_on arr_to_hfun (hom (hf_to_ide A) (hf_to_ide C))"
proof -
have "ide (hf_to_ide A) \<and> ide (hf_to_ide C)"
using assms hf_to_ide_mapsto by auto
thus ?thesis
using bij_betw_hom_hfun [of "hf_to_ide A" "hf_to_ide C"] bij_betw_imp_inj_on
by auto
qed
ultimately show ?thesis
using 1 inj_on_def [of arr_to_hfun "hom (hf_to_ide A) (hf_to_ide C)"] by simp
qed
qed
section "Binary Products"
text\<open>
The category of hereditarily finite sets has binary products,
given by cartesian product of sets in the usual way.
\<close>
definition prod
where "prod a b = hf_to_ide (ide_to_hf a * ide_to_hf b)"
definition pr0
where "pr0 a b = (if ide a \<and> ide b then
mkArr (set (prod a b)) (set b) (\<lambda>x. arr_of (hsnd (elem_of x)))
else null)"
definition pr1
where "pr1 a b = (if ide a \<and> ide b then
mkArr (set (prod a b)) (set a) (\<lambda>x. arr_of (hfst (elem_of x)))
else null)"
definition tuple
where "tuple f g = mkArr (set (dom f)) (set (prod (cod f) (cod g)))
(\<lambda>x. arr_of (hpair (elem_of (Fun f x)) (elem_of (Fun g x))))"
lemma ide_prod:
assumes "ide a" and "ide b"
shows "ide (prod a b)"
using assms prod_def hf_to_ide_mapsto ide_to_hf_mapsto by auto
lemma pr1_in_hom [intro]:
assumes "ide a" and "ide b"
shows "\<guillemotleft>pr1 a b : prod a b \<rightarrow> a\<guillemotright>"
proof
show 0: "arr (pr1 a b)"
proof -
have "set (prod a b) \<subseteq> Univ \<and> finite (set (prod a b))"
using assms ide_implies_finite_set(1) set_ideD(1) ide_prod by presburger
moreover have "set a \<subseteq> Univ \<and> finite (set a)"
using assms ide_char set_ide_char by blast
moreover have "(\<lambda>x. arr_of (hfst (elem_of x))) \<in> set (prod a b) \<rightarrow> set a"
proof (unfold prod_def)
show "(\<lambda>x. arr_of (hfst (elem_of x))) \<in> set (hf_to_ide (ide_to_hf a * ide_to_hf b)) \<rightarrow> set a"
proof
fix x
assume x: "x \<in> set (hf_to_ide (ide_to_hf a * ide_to_hf b))"
have "elem_of x \<in> hfset (ide_to_hf a * ide_to_hf b)"
using assms ide_char x
by (metis (no_types, lifting) prod_def elem_of_membI HF_hfset UNIV_I hmem_HF_iff
ide_prod ide_to_hf_hf_to_ide)
hence "hfst (elem_of x) \<^bold>\<in> ide_to_hf a"
by (metis HF_hfset finite_hfset hfst_conv hmem_HF_iff timesE)
thus "arr_of (hfst (elem_of x)) \<in> set a"
using arr_of_membI by simp
qed
qed
ultimately show ?thesis
unfolding pr1_def
using assms arr_mkArr finite_imp_setp by presburger
qed
show "dom (pr1 a b) = prod a b"
using assms 0 ide_char ide_prod dom_mkArr
by (metis (no_types, lifting) mkIde_set pr1_def)
show "cod (pr1 a b) = a"
using assms 0 ide_char ide_prod cod_mkArr
by (metis (no_types, lifting) mkIde_set pr1_def)
qed
lemma pr1_simps [simp]:
assumes "ide a" and "ide b"
shows "arr (pr1 a b)" and "dom (pr1 a b) = prod a b" and "cod (pr1 a b) = a"
using assms pr1_in_hom by blast+
lemma pr0_in_hom [intro]:
assumes "ide a" and "ide b"
shows "\<guillemotleft>pr0 a b : prod a b \<rightarrow> b\<guillemotright>"
proof
show 0: "arr (pr0 a b)"
proof -
have "set (prod a b) \<subseteq> Univ \<and> finite (set (prod a b))"
using setp_set_ide assms ide_implies_finite_set(1) ide_prod by presburger
moreover have "set b \<subseteq> Univ \<and> finite (set b)"
using assms ide_char set_ide_char by blast
moreover have "(\<lambda>x. arr_of (hsnd (elem_of x))) \<in> set (prod a b) \<rightarrow> set b"
proof (unfold prod_def)
show "(\<lambda>x. arr_of (hsnd (elem_of x))) \<in> set (hf_to_ide (ide_to_hf a * ide_to_hf b)) \<rightarrow> set b"
proof
fix x
assume x: "x \<in> set (hf_to_ide (ide_to_hf a * ide_to_hf b))"
have "elem_of x \<in> hfset (ide_to_hf a * ide_to_hf b)"
using assms ide_char x
by (metis (no_types, lifting) prod_def elem_of_membI HF_hfset UNIV_I hmem_HF_iff
ide_prod ide_to_hf_hf_to_ide)
hence "hsnd (elem_of x) \<^bold>\<in> ide_to_hf b"
by (metis HF_hfset finite_hfset hsnd_conv hmem_HF_iff timesE)
thus "arr_of (hsnd (elem_of x)) \<in> set b"
using arr_of_membI by simp
qed
qed
ultimately show ?thesis
unfolding pr0_def
using assms arr_mkArr finite_imp_setp by presburger
qed
show "dom (pr0 a b) = prod a b"
using assms 0 ide_char ide_prod dom_mkArr
by (metis (no_types, lifting) mkIde_set pr0_def)
show "cod (pr0 a b) = b"
using assms 0 ide_char ide_prod cod_mkArr
by (metis (no_types, lifting) mkIde_set pr0_def)
qed
lemma pr0_simps [simp]:
assumes "ide a" and "ide b"
shows "arr (pr0 a b)" and "dom (pr0 a b) = prod a b" and "cod (pr0 a b) = b"
using assms pr0_in_hom by blast+
lemma arr_of_tuple_elem_of_membI:
assumes "span f g" and "x \<in> Dom f"
shows "arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle> \<in> set (prod (cod f) (cod g))"
proof -
have "Fun f x \<in> set (cod f)"
using assms Fun_mapsto by blast
moreover have "Fun g x \<in> set (cod g)"
using assms Fun_mapsto by auto
ultimately have "\<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>
\<^bold>\<in> ide_to_hf (cod f) * ide_to_hf (cod g)"
using assms ide_cod by auto
moreover have "set (prod (cod f) (cod g)) \<subseteq> Univ"
using setp_set_ide assms(1) ide_cod ide_prod by presburger
ultimately show ?thesis
using prod_def arr_of_membI ide_to_hf_hf_to_ide by auto
qed
lemma tuple_in_hom [intro]:
assumes "span f g"
shows "\<guillemotleft>tuple f g : dom f \<rightarrow> prod (cod f) (cod g)\<guillemotright>"
proof
show 1: "arr (tuple f g)"
proof -
have "Dom f \<subseteq> Univ \<and> finite (Dom f)"
using assms set_ideD(1) ide_dom ide_implies_finite_set(1) by presburger
moreover have "set (prod (cod f) (cod g)) \<subseteq> Univ \<and> finite (set (prod (cod f) (cod g)))"
using assms set_ideD(1) ide_cod ide_prod ide_implies_finite_set(1) by presburger
moreover have "(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>)
\<in> Dom f \<rightarrow> set (prod (cod f) (cod g))"
using assms arr_of_tuple_elem_of_membI by simp
ultimately show ?thesis
using assms ide_prod tuple_def arr_mkArr ide_dom ide_cod by simp
qed
show "dom (tuple f g) = dom f"
using assms 1 dom_mkArr ide_dom mkIde_set tuple_def by auto
show "cod (tuple f g) = prod (cod f) (cod g)"
using assms 1 cod_mkArr ide_cod mkIde_set tuple_def ide_prod by auto
qed
lemma tuple_simps [simp]:
assumes "span f g"
shows "arr (tuple f g)" and "dom (tuple f g) = dom f"
and "cod (tuple f g) = prod (cod f) (cod g)"
using assms tuple_in_hom by blast+
lemma Fun_pr1:
assumes "ide a" and "ide b"
shows "Fun (pr1 a b) = restrict (\<lambda>x. arr_of (hfst (elem_of x))) (set (prod a b))"
using assms pr1_def Fun_mkArr arr_char pr1_simps(1) by presburger
lemma Fun_pr0:
assumes "ide a" and "ide b"
shows "Fun (pr0 a b) = restrict (\<lambda>x. arr_of (hsnd (elem_of x))) (set (prod a b))"
using assms pr0_def Fun_mkArr arr_char pr0_simps(1) by presburger
lemma Fun_tuple:
assumes "span f g"
shows "Fun (tuple f g) = restrict (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>) (Dom f)"
proof -
have "arr (tuple f g)"
using assms tuple_in_hom by blast
thus ?thesis
using assms tuple_def Fun_mkArr by simp
qed
lemma pr1_tuple:
assumes "span f g"
shows "comp (pr1 (cod f) (cod g)) (tuple f g) = f"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have pr1: "\<guillemotleft>pr1 (cod f) (cod g) : prod (cod f) (cod g) \<rightarrow> cod f\<guillemotright>"
using assms ide_cod by blast
have tuple: "\<guillemotleft>tuple f g : dom f \<rightarrow> prod (cod f) (cod g)\<guillemotright>"
using assms by blast
show par: "par (comp (pr1 (cod f) (cod g)) (tuple f g)) f"
using assms pr1_in_hom tuple_in_hom
by (metis (no_types, lifting) comp_in_homI' ide_cod in_homE)
show "Fun (comp (pr1 (cod f) (cod g)) (tuple f g)) = Fun f"
proof -
have seq: "seq (pr1 (cod f) (cod g)) (tuple f g)"
using par by blast
have "Fun (comp (pr1 (cod f) (cod g)) (tuple f g)) =
restrict (Fun (pr1 (cod f) (cod g)) \<circ> Fun (tuple f g)) (Dom (tuple f g))"
using pr1 tuple seq Fun_comp by simp
also have "... = restrict
(Fun (mkArr (set (prod (cod f) (cod g))) (Cod f)
(\<lambda>x. arr_of (hfst (elem_of x)))) \<circ>
Fun (mkArr (Dom f) (set (prod (cod f) (cod g)))
(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>)))
(Dom (tuple f g))"
unfolding pr1_def tuple_def
using assms ide_cod by presburger
also have
"... = restrict
(restrict (\<lambda>x. arr_of (hfst (elem_of x))) (set (prod (cod f) (cod g))) o
restrict (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>) (Dom f))
(Dom f)"
proof -
have "Fun (mkArr (set (prod (cod f) (cod g))) (Cod f) (\<lambda>x. arr_of (hfst (elem_of x)))) =
restrict (\<lambda>x. arr_of (hfst (elem_of x))) (set (prod (cod f) (cod g)))"
using assms Fun_mkArr ide_prod pr1
by (metis (no_types, lifting) arrI ide_cod pr1_def)
moreover have "Fun (mkArr (Dom f) (set (prod (cod f) (cod g)))
(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>)) =
restrict (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>) (Dom f)"
using assms Fun_mkArr ide_prod ide_cod tuple_def tuple arrI by simp
ultimately show ?thesis
using assms tuple_simps(2) by simp
qed
also have
"... = restrict
((\<lambda>x. arr_of (hfst (elem_of x))) o (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>))
(Dom f)"
using assms tuple tuple_def arr_of_tuple_elem_of_membI by auto
also have "... = restrict (Fun f) (Dom f)"
proof
fix x
have "restrict ((\<lambda>x. arr_of (hfst (elem_of x))) o (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>))
(Dom f) x =
restrict (\<lambda>x. arr_of (elem_of (Fun f x))) (Dom f) x"
by simp
also have "... = restrict (Fun f) (Dom f) x"
proof (cases "x \<in> Dom f")
show "x \<notin> Dom f \<Longrightarrow> ?thesis" by simp
assume x: "x \<in> Dom f"
have "Fun f x \<in> Cod f"
using assms x Fun_mapsto arr_char by blast
moreover have "Cod f \<subseteq> Univ"
using setp_set_ide assms ide_cod by blast
ultimately show ?thesis
using assms arr_of_elem_of Fun_mapsto by auto
qed
finally show "restrict ((\<lambda>x. arr_of (hfst (elem_of x))) \<circ>
(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>))
(Dom f) x =
restrict (Fun f) (Dom f) x"
by blast
qed
also have "... = Fun f"
using assms par Fun_mapsto Fun_mkArr mkArr_Fun
by (metis (no_types, lifting))
finally show ?thesis by blast
qed
qed
lemma pr0_tuple:
assumes "span f g"
shows "comp (pr0 (cod f) (cod g)) (tuple f g) = g"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have pr0: "\<guillemotleft>pr0 (cod f) (cod g) : prod (cod f) (cod g) \<rightarrow> cod g\<guillemotright>"
using assms ide_cod by blast
have tuple: "\<guillemotleft>tuple f g : dom f \<rightarrow> prod (cod f) (cod g)\<guillemotright>"
using assms by blast
show par: "par (comp (pr0 (cod f) (cod g)) (tuple f g)) g"
using assms pr0_in_hom tuple_in_hom
by (metis (no_types, lifting) comp_in_homI' ide_cod in_homE)
show "Fun (comp (pr0 (cod f) (cod g)) (tuple f g)) = Fun g"
proof -
have seq: "seq (pr0 (cod f) (cod g)) (tuple f g)"
using par by blast
have "Fun (comp (pr0 (cod f) (cod g)) (tuple f g)) =
restrict (Fun (pr0 (cod f) (cod g)) \<circ> Fun (tuple f g)) (Dom (tuple f g))"
using pr0 tuple seq Fun_comp by simp
also have
"... = restrict
(Fun (mkArr (set (prod (cod f) (cod g))) (Cod g)
(\<lambda>x. arr_of (hsnd (elem_of x)))) \<circ>
Fun (mkArr (Dom f) (set (prod (cod f) (cod g)))
(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>)))
(Dom (tuple f g))"
unfolding pr0_def tuple_def
using assms ide_cod by presburger
also have "... = restrict
(restrict (\<lambda>x. arr_of (hsnd (elem_of x))) (set (prod (cod f) (cod g))) o
restrict (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>) (Dom g))
(Dom g)"
proof -
have "Fun (mkArr (set (prod (cod f) (cod g))) (Cod g) (\<lambda>x. arr_of (hsnd (elem_of x)))) =
restrict (\<lambda>x. arr_of (hsnd (elem_of x))) (set (prod (cod f) (cod g)))"
using assms Fun_mkArr ide_prod arrI
by (metis (no_types, lifting) ide_cod pr0 pr0_def)
moreover have "Fun (mkArr (Dom f) (set (prod (cod f) (cod g)))
(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>)) =
restrict (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>) (Dom f)"
using assms Fun_mkArr ide_prod ide_cod tuple_def tuple arrI by simp
ultimately show ?thesis
using assms tuple_simps(2) by simp
qed
also have "... = restrict
((\<lambda>x. arr_of (hsnd (elem_of x))) o (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>))
(Dom g)"
using assms tuple tuple_def arr_of_tuple_elem_of_membI by auto
also have "... = restrict (Fun g) (Dom g)"
proof
fix x
have "restrict ((\<lambda>x. arr_of (hsnd (elem_of x)))
o (\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>))
(Dom g) x =
restrict (\<lambda>x. arr_of (elem_of (Fun g x))) (Dom g) x"
by simp
also have "... = restrict (Fun g) (Dom g) x"
proof (cases "x \<in> Dom g")
show "x \<notin> Dom g \<Longrightarrow> ?thesis" by simp
assume x: "x \<in> Dom g"
have "Fun g x \<in> Cod g"
using assms x Fun_mapsto arr_char by blast
moreover have "Cod g \<subseteq> Univ"
using assms set_ideD(1) ide_cod by blast
ultimately show ?thesis
using assms arr_of_elem_of Fun_mapsto by auto
qed
finally show "restrict ((\<lambda>x. arr_of (hsnd (elem_of x))) \<circ>
(\<lambda>x. arr_of \<langle>elem_of (Fun f x), elem_of (Fun g x)\<rangle>))
(Dom g) x =
restrict (Fun g) (Dom g) x"
by blast
qed
also have "... = Fun g"
using assms par Fun_mapsto Fun_mkArr mkArr_Fun
by (metis (no_types, lifting))
finally show ?thesis by blast
qed
qed
lemma tuple_pr:
assumes "ide a" and "ide b" and "\<guillemotleft>h : dom h \<rightarrow> prod a b\<guillemotright>"
shows "tuple (comp (pr1 a b) h) (comp (pr0 a b) h) = h"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have pr0: "\<guillemotleft>pr0 a b : prod a b \<rightarrow> b\<guillemotright>"
using assms pr0_in_hom ide_cod by blast
have pr1: "\<guillemotleft>pr1 a b : prod a b \<rightarrow> a\<guillemotright>"
using assms pr1_in_hom ide_cod by blast
have tuple: "\<guillemotleft>tuple (comp (pr1 a b) h) (comp (pr0 a b) h) : dom h \<rightarrow> prod a b\<guillemotright>"
using assms pr0 pr1
by (metis (no_types, lifting) cod_comp dom_comp pr0_simps(3) pr1_simps(3)
seqI' tuple_in_hom)
show par: "par (tuple (comp (pr1 a b) h) (comp (pr0 a b) h)) h"
using assms tuple by (metis (no_types, lifting) in_homE)
show "Fun (tuple (comp (pr1 a b) h) (comp (pr0 a b) h)) = Fun h"
proof -
have 1: "Fun (comp (pr1 a b) h) =
restrict (restrict (\<lambda>x. arr_of (hfst (elem_of x))) (set (prod a b)) \<circ> Fun h) (Dom h)"
using assms pr1 Fun_comp Fun_pr1 seqI' by auto
have 2: "Fun (comp (pr0 a b) h) =
restrict (restrict (\<lambda>x. arr_of (hsnd (elem_of x))) (set (prod a b)) \<circ> Fun h) (Dom h)"
using assms pr0 Fun_comp Fun_pr0 seqI' by auto
have "Fun (tuple (comp (pr1 a b) h) (comp (pr0 a b) h)) =
restrict (\<lambda>x. arr_of \<langle>elem_of (restrict
(restrict (\<lambda>x. arr_of (hfst (elem_of x))) (set (prod a b)) \<circ> Fun h)
(Dom h) x),
elem_of (restrict
(restrict (\<lambda>x. arr_of (hsnd (elem_of x))) (set (prod a b)) \<circ> Fun h)
(Dom h) x)\<rangle>)
(Dom h)"
proof -
have "Dom (comp (pr1 a b) h) = Dom h"
using assms pr1_in_hom
by (metis (no_types, lifting) in_homE dom_comp seqI)
moreover have "arr (mkArr (Dom (comp (pr1 a b) h))
(set (prod (cod (comp (pr1 a b) h)) (cod (comp (pr0 a b) h))))
(\<lambda>x. arr_of \<langle>elem_of (Fun (comp (pr1 a b) h) x),
elem_of (Fun (comp (pr0 a b) h) x)\<rangle>))"
using tuple unfolding tuple_def by blast
ultimately show ?thesis
using 1 2 tuple tuple_def
Fun_mkArr [of "Dom (comp (pr1 a b) h)"
"set (prod (cod (comp (pr1 a b) h))
(cod (comp (pr0 a b) h)))"
"\<lambda>x. arr_of \<langle>elem_of (Fun (comp (pr1 a b) h) x),
elem_of (Fun (comp (pr0 a b) h) x)\<rangle>"]
by simp
qed
also have "... = Fun h"
proof
let ?f = "..."
fix x
show "?f x = Fun h x"
proof -
have "x \<notin> Dom h \<Longrightarrow> ?f x = Fun h x"
proof -
assume x: "x \<notin> Dom h"
have "restrict ?f (Dom h) x = undefined"
using assms x restrict_apply by auto
also have "... = Fun h x"
proof -
have "arr h"
using assms by blast
thus ?thesis
using assms x Fun_mapsto [of h] extensional_arb [of "Fun h" "Dom h" x]
by simp
qed
finally show ?thesis by auto
qed
moreover have "x \<in> Dom h \<Longrightarrow> ?f x = Fun h x"
proof -
assume x: "x \<in> Dom h"
have 1: "Fun h x \<in> set (prod a b)"
proof -
have "Fun h x \<in> Cod h"
using assms x Fun_mapsto [of h] by blast
moreover have "Cod h = set (prod a b)"
using assms ide_prod
by (metis (no_types, lifting) in_homE)
ultimately show ?thesis by fast
qed
have "?f x = arr_of \<langle>hfst (elem_of (Fun h x)), hsnd (elem_of (Fun h x))\<rangle>"
using x 1 by simp
also have "... = arr_of (elem_of (Fun h x))"
proof -
have "elem_of (Fun h x) \<^bold>\<in> ide_to_hf a * ide_to_hf b"
using assms x 1 par
by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I ide_prod
ide_to_hf_hf_to_ide)
thus ?thesis
using x is_hpair_def by auto
qed
also have "... = Fun h x"
using 1 arr_of_elem_of assms set_ideD(1) ide_prod by blast
finally show ?thesis by blast
qed
ultimately show ?thesis by blast
qed
qed
finally show ?thesis by blast
qed
qed
interpretation HF': elementary_category_with_binary_products comp pr0 pr1
proof
show "\<And>a b. \<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> span (pr1 a b) (pr0 a b)"
using pr0_simps(1) pr0_simps(2) pr1_simps(1) pr1_simps(2) by auto
show "\<And>a b. \<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> cod (pr0 a b) = b"
using pr0_simps(1-3) by blast
show "\<And>a b. \<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> cod (pr1 a b) = a"
using pr1_simps(1-3) by blast
show "\<And>f g. span f g \<Longrightarrow>
\<exists>!l. comp (pr1 (cod f) (cod g)) l = f \<and> comp (pr0 (cod f) (cod g)) l = g"
proof
fix f g
assume fg: "span f g"
show "comp (pr1 (cod f) (cod g)) (tuple f g) = f \<and>
comp (pr0 (cod f) (cod g)) (tuple f g) = g"
using fg pr0_simps pr1_simps tuple_simps pr0_tuple pr1_tuple by presburger
show "\<And>l. \<lbrakk>comp (pr1 (cod f) (cod g)) l = f \<and> comp (pr0 (cod f) (cod g)) l = g\<rbrakk>
\<Longrightarrow> l = tuple f g "
proof -
fix l
assume l: "comp (pr1 (cod f) (cod g)) l = f \<and> comp (pr0 (cod f) (cod g)) l = g"
show "l = tuple f g"
using fg l tuple_pr
by (metis (no_types, lifting) arr_iff_in_hom ide_cod seqE pr0_simps(2))
qed
qed
show "\<And>a b. \<not> (ide a \<and> ide b) \<Longrightarrow> pr0 a b = null"
using pr0_def by auto
show "\<And>a b. \<not> (ide a \<and> ide b) \<Longrightarrow> pr1 a b = null"
using pr1_def by auto
qed
text\<open>
For reasons of economy of locale parameters, the notion \<open>prod\<close> is a defined notion
of the @{locale elementary_category_with_binary_products} locale.
However, we need to be able to relate this notion to that of cartesian product of
hereditarily finite sets, which we have already used to give a definition of \<open>prod\<close>.
The locale assumptions for @{locale elementary_cartesian_closed_category} refer
specifically to \<open>HF'.prod\<close>, even though in the end the notion itself does not depend
on that choice. To be able to show that the locale assumptions of
@{locale elementary_cartesian_closed_category} are satisfied, we need to use a choice
of products that we can relate to the cartesian product of hereditarily
finite sets. We therefore need to show that our previously defined \<open>prod\<close> coincides
(on objects) with the one defined in the @{locale elementary_category_with_binary_products} locale;
\emph{i.e.}~\<open>HF'.prod\<close>. Note that the latter is defined for all arrows,
not just identity arrows, so we need to use that for the subsequent definitions and proofs.
\<close>
lemma prod_ide_eq:
assumes "ide a" and "ide b"
shows "prod a b = HF'.prod a b"
using assms prod_def HF'.pr_simps(2) HF'.prod_def pr0_simps(2) by presburger
lemma tuple_span_eq:
assumes "span f g"
shows "tuple f g = HF'.tuple f g"
using assms tuple_def HF'.tuple_def
by (metis (no_types, lifting) HF'.tuple_eqI pr0_tuple pr1_tuple)
section "Exponentials"
text\<open>
We now turn our attention to exponentials.
\<close>
definition exp
where "exp b c = hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c))"
definition eval
where "eval b c = mkArr (set (HF'.prod (exp b c) b)) (set c)
(\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))"
definition \<Lambda>
where "\<Lambda> a b c f = mkArr (set a) (set (exp b c))
(\<lambda>x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c)
(arr_to_hfun f))
(elem_of x)))"
lemma ide_exp:
assumes "ide b" and "ide c"
shows "ide (exp b c)"
using assms exp_def hf_to_ide_mapsto ide_to_hf_mapsto by auto
lemma hfset_ide_to_hf:
assumes "ide a"
shows "hfset (ide_to_hf a) = elem_of ` set a"
using assms ide_to_hf_def ide_implies_finite_set(1) by auto
lemma eval_in_hom [intro]:
assumes "ide b" and "ide c"
shows "in_hom (eval b c) (HF'.prod (exp b c) b) c"
proof
show 1: "arr (eval b c)"
proof (unfold eval_def arr_mkArr, intro conjI)
show "set (HF'.prod (exp b c) b) \<subseteq> Univ"
using HF'.ide_prod assms set_ideD(1) ide_exp by presburger
show "set c \<subseteq> Univ"
using assms set_ideD(1) by blast
show "(\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
\<in> set (HF'.prod (exp b c) b) \<rightarrow> set c"
proof
fix x
assume "x \<in> set (HF'.prod (exp b c) b)"
hence x: "x \<in> set (prod (exp b c) b)"
using assms prod_ide_eq ide_exp by auto
show "arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))) \<in> set c"
proof (intro arr_of_membI)
show "happ (hfst (elem_of x)) (hsnd (elem_of x)) \<^bold>\<in> ide_to_hf c"
proof -
have 1: "elem_of x \<^bold>\<in> ide_to_hf (exp b c) * ide_to_hf b"
proof -
have "elem_of x \<^bold>\<in> ide_to_hf (prod (exp b c) b)"
using assms x elem_of_membI ide_prod ide_exp by simp
thus ?thesis
using assms x prod_def ide_to_hf_hf_to_ide by auto
qed
have "hfst (elem_of x) \<^bold>\<in> hexp (ide_to_hf b) (ide_to_hf c)"
using assms 1 x exp_def ide_to_hf_hf_to_ide by auto
moreover have "hsnd (elem_of x) \<^bold>\<in> ide_to_hf b"
using assms 1 by auto
ultimately show ?thesis
using happ_mapsto [of "hfst (elem_of x)" "ide_to_hf b" "ide_to_hf c"
"hsnd (elem_of x)"]
by simp
qed
qed
qed
show "finite (elem_of ` set (HF'.prod (exp b c) b))"
using HF'.ide_prod setp_set_ide assms ide_exp by presburger
show "finite (elem_of ` set c)"
using setp_set_ide assms(2) by blast
qed
show "dom (eval b c) = HF'.prod (exp b c) b"
using assms 1 ide_char HF'.ide_prod ide_exp dom_mkArr eval_def
by (metis (no_types, lifting) mkIde_set)
show "cod (eval b c) = c"
using assms 1 ide_char cod_mkArr eval_def
by (metis (no_types, lifting) mkIde_set)
qed
lemma eval_simps [simp]:
assumes "ide b" and "ide c"
shows "arr (eval b c)"
and "dom (eval b c) = HF'.prod (exp b c) b"
and "cod (eval b c) = c"
using assms eval_in_hom by blast+
lemma hlam_arr_to_hfun_in_hexp:
assumes "ide a" and "ide b" and "ide c"
and "in_hom f (prod a b) c"
shows "hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)
\<^bold>\<in> hexp (ide_to_hf a) (ide_to_hf (exp b c))"
using assms hfun_in_hexp hfun_hlam
by (metis (no_types, lifting) prod_def HCollect_iff in_homE UNIV_I
arr_to_hfun_in_hexp exp_def hexp_def ide_to_hf_hf_to_ide)
lemma lam_in_hom [intro]:
assumes "ide a" and "ide b" and "ide c"
and "in_hom f (prod a b) c"
shows "in_hom (\<Lambda> a b c f) a (exp b c)"
proof
show 1: "arr (\<Lambda> a b c f)"
proof (unfold \<Lambda>_def arr_mkArr, intro conjI)
show "set a \<subseteq> Univ"
using assms(1) set_ideD(1) by blast
show "set (exp b c) \<subseteq> Univ"
using assms(2-3) set_ideD(1) ide_exp ide_char by blast
show "finite (elem_of ` set a)"
using assms(1) set_ideD(1) setp_set_ide by presburger
show "finite (elem_of ` set (exp b c))"
using assms(2-3) set_ideD(1) setp_set_ide ide_exp by presburger
show "(\<lambda>x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f))
(elem_of x)))
\<in> set a \<rightarrow> set (exp b c)"
proof
fix x
assume x: "x \<in> set a"
show "arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f))
(elem_of x))
\<in> set (exp b c)"
using assms x hlam_arr_to_hfun_in_hexp ide_to_hf_def elem_of_membI happ_mapsto
arr_of_membI
by meson
qed
qed
show "dom (\<Lambda> a b c f) = a"
using assms(1) 1 \<Lambda>_def ide_char dom_mkArr mkIde_set by auto
show "cod (\<Lambda> a b c f) = exp b c"
using assms(2-3) 1 \<Lambda>_def cod_mkArr ide_exp mkIde_set by auto
qed
lemma lam_simps [simp]:
assumes "ide a" and "ide b" and "ide c"
and "in_hom f (prod a b) c"
shows "arr (\<Lambda> a b c f)"
and "dom (\<Lambda> a b c f) = a"
and "cod (\<Lambda> a b c f) = exp b c"
using assms lam_in_hom by blast+
lemma Fun_lam:
assumes "ide a" and "ide b" and "ide c"
and "in_hom f (prod a b) c"
shows "Fun (\<Lambda> a b c f) =
restrict (\<lambda>x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f))
(elem_of x)))
(set a)"
using assms arr_char lam_simps(1) \<Lambda>_def Fun_mkArr by simp
lemma Fun_eval:
assumes "ide b" and "ide c"
shows "Fun (eval b c) = restrict (\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
(set (HF'.prod (exp b c) b))"
using assms arr_char eval_simps(1) eval_def Fun_mkArr by force
lemma Fun_prod:
assumes "arr f" and "arr g" and "x \<in> set (prod (dom f) (dom g))"
shows "Fun (HF'.prod f g) x = arr_of \<langle>elem_of (Fun f (arr_of (hfst (elem_of x)))),
elem_of (Fun g (arr_of (hsnd (elem_of x))))\<rangle>"
proof -
have 1: "span (comp f (pr1 (dom f) (dom g))) (comp g (pr0 (dom f) (dom g)))"
using assms
by (metis (no_types, lifting) HF'.prod_def HF'.prod_simps(1) HF'.tuple_ext not_arr_null)
have 2: "Dom (comp f (pr1 (dom f) (dom g))) = set (prod (dom f) (dom g))"
using assms
by (metis (mono_tags, lifting) 1 dom_comp ide_dom pr0_simps(2))
have 3: "Dom (comp g (pr0 (dom f) (dom g))) = set (prod (dom f) (dom g))"
using assms 1 2 by force
have "Fun (HF'.prod f g) x =
Fun (HF'.tuple (comp f (pr1 (dom f) (dom g))) (comp g (pr0 (dom f) (dom g)))) x"
using assms(3) HF'.prod_def by simp
also have "... = restrict (\<lambda>x. arr_of \<langle>elem_of (Fun (comp f (pr1 (dom f) (dom g))) x),
elem_of (Fun (comp g (pr0 (dom f) (dom g))) x)\<rangle>)
(Dom (comp f (pr1 (dom f) (dom g))))
x"
using assms 1 tuple_span_eq Fun_tuple by simp
also have "... = arr_of \<langle>elem_of (Fun (comp f (pr1 (dom f) (dom g))) x),
elem_of (Fun (comp g (pr0 (dom f) (dom g))) x)\<rangle>"
using assms(3) 2 by simp
also have "... = arr_of \<langle>elem_of (Fun f (arr_of (hfst (elem_of x)))),
elem_of (Fun g (arr_of (hsnd (elem_of x))))\<rangle>"
proof -
have "Fun (comp f (pr1 (dom f) (dom g))) x = Fun f (arr_of (hfst (elem_of x)))"
proof -
(* TODO: Figure out what is making this proof so "stiff". *)
have 4: "seq f (pr1 (dom f) (dom g))"
using assms 1 by blast
have "Fun (comp f (pr1 (dom f) (dom g))) x =
restrict (Fun f \<circ> Fun (pr1 (dom f) (dom g))) (Dom (pr1 (dom f) (dom g))) x"
using assms 1 Fun_comp [of f "pr1 (dom f) (dom g)"]
by (metis (no_types, lifting))
also have "... = (Fun f \<circ> Fun (pr1 (dom f) (dom g))) x"
proof -
have "x \<in> Dom (pr1 (dom f) (dom g))"
using assms 1 2 4
by (metis (no_types, lifting) dom_comp)
thus ?thesis by simp
qed
also have "... = Fun f (Fun (pr1 (dom f) (dom g)) x)"
by simp
also have "... = Fun f (arr_of (hfst (elem_of x)))"
using assms 1 Fun_pr1 [of "dom f" "dom g"] ide_dom by simp
finally show ?thesis by blast
qed
moreover
have "Fun (comp g (pr0 (dom f) (dom g))) x = Fun g (arr_of (hsnd (elem_of x)))"
proof -
have 4: "seq g (pr0 (dom f) (dom g))"
using assms 1 by blast
have "Fun (comp g (pr0 (dom f) (dom g))) x =
restrict (Fun g \<circ> Fun (pr0 (dom f) (dom g))) (Dom (pr0 (dom f) (dom g))) x"
using assms 1 Fun_comp [of g "pr0 (dom f) (dom g)"]
by (metis (no_types, lifting))
also have "... = (Fun g \<circ> Fun (pr0 (dom f) (dom g))) x"
proof -
have "x \<in> Dom (pr0 (dom f) (dom g))"
using assms 1 2 4
by (metis (no_types, lifting) dom_comp)
thus ?thesis by simp
qed
also have "... = Fun g (Fun (pr0 (dom f) (dom g)) x)"
by simp
also have "... = Fun g (arr_of (hsnd (elem_of x)))"
using assms 1 Fun_pr0 [of "dom f" "dom g"] ide_dom by simp
finally show ?thesis by blast
qed
ultimately show ?thesis by simp
qed
finally show ?thesis by simp
qed
lemma prod_in_terms_of_tuple:
assumes "arr f" and "arr g"
shows "HF'.prod f g =
tuple (comp f (pr1 (dom f) (dom g))) (comp g (pr0 (dom f) (dom g)))"
using assms HF'.prod_def tuple_span_eq
by (metis (no_types, lifting) HF'.prod_simps(1) HF'.tuple_ext not_arr_null)
lemma eval_prod_lam:
assumes "ide a" and "ide b" and "ide c"
and "in_hom g (prod a b) c"
shows "comp (eval b c) (HF'.prod (\<Lambda> a b c g) b) = g"
proof -
have ide_dom_lam: "ide (dom (\<Lambda> a b c g))"
using assms lam_in_hom [of a b c g] ide_dom by blast
have ide_dom_b: "ide (dom b)"
using assms ide_dom ideD(1) by blast
define \<Lambda>_pr1 where "\<Lambda>_pr1 = comp (\<Lambda> a b c g) (pr1 (dom (\<Lambda> a b c g)) (dom b))"
define b_pr0 where "b_pr0 = comp b (pr0 (dom (\<Lambda> a b c g)) (dom b))"
have lam_pr1: "in_hom \<Lambda>_pr1 (prod a b) (exp b c)"
proof (unfold \<Lambda>_pr1_def, intro comp_in_homI)
show "in_hom (pr1 (dom (\<Lambda> a b c g)) (dom b)) (prod a b) a"
using assms ide_dom_lam ide_dom_b ideD(2) lam_simps(2) pr1_in_hom by auto
show "in_hom (\<Lambda> a b c g) a (exp b c)"
using assms lam_in_hom by simp
qed
have b_pr0: "in_hom b_pr0 (prod a b) b"
using assms b_pr0_def
by (metis (no_types, lifting) HF'.arr_pr0_iff HF'.cod_pr0 comp_in_homI'
ideD(1-3) lam_simps(2) pr0_simps(2))
have 1: "span \<Lambda>_pr1 b_pr0"
using lam_pr1 b_pr0
by (metis (no_types, lifting) in_homE)
have tuple: "in_hom (tuple \<Lambda>_pr1 b_pr0) (prod a b) (prod (exp b c) b)"
using 1 lam_pr1 b_pr0 tuple_in_hom [of \<Lambda>_pr1 b_pr0]
by (metis (mono_tags, lifting) in_homE)
define \<Lambda>_pr1' where "\<Lambda>_pr1' = comp (\<Lambda> a b c g) (pr1 a b)"
define b_pr0' where "b_pr0' = pr0 a b"
have lam_pr1_eq: "\<Lambda>_pr1 = \<Lambda>_pr1'"
using assms \<Lambda>_pr1_def \<Lambda>_pr1'_def ideD(2) lam_simps(2) by auto
have b_pr0_eq: "b_pr0 = b_pr0'"
using assms b_pr0_def b_pr0'_def b_pr0 comp_ide_arr
by (metis (no_types, lifting) ideD(2) in_homE lam_simps(2))
have Fun_pr0: "Fun (pr0 a b) = restrict (\<lambda>x. arr_of (hsnd (elem_of x))) (set (prod a b))"
using assms Fun_pr0 by simp
have Fun_lam_pr1: "Fun \<Lambda>_pr1 =
restrict (Fun (\<Lambda> a b c g) o
restrict (\<lambda>x. arr_of (hfst (elem_of x))) (set (prod a b)))
(set (prod a b))"
using assms 1 Fun_comp Fun_pr1 lam_pr1_eq \<Lambda>_pr1'_def
by (metis (no_types, lifting) pr1_simps(2))
have "comp (eval b c) (HF'.prod (\<Lambda> a b c g) b) = comp (eval b c) (tuple \<Lambda>_pr1 b_pr0)"
using assms \<Lambda>_pr1_def b_pr0_def 1 prod_in_terms_of_tuple ideD(1) lam_simps(1)
by presburger
also have 5: "... = comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')"
using lam_pr1_eq b_pr0_eq by simp
also have "... = g"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 2: "arr (comp (eval b c) (tuple \<Lambda>_pr1 b_pr0))"
using assms tuple arr_char
by (metis (no_types, lifting) in_homE seqI eval_simps(1-2) ide_exp prod_ide_eq)
have 3: "arr g"
using assms by blast
have tuple': "in_hom (tuple \<Lambda>_pr1' b_pr0') (prod a b) (prod (exp b c) b)"
using tuple lam_pr1_eq b_pr0_eq by blast
have 4: "Dom g = set (prod a b)"
using assms
by (metis (no_types, lifting) in_homE)
show par: "par (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) g"
using assms tuple' 2 3 5
by (metis (no_types, lifting) cod_comp dom_comp in_homE eval_simps(3))
show "Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) = Fun g"
proof
fix x
have "x \<notin> set (prod a b) \<Longrightarrow> Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) x = Fun g x"
proof -
have 5: "Fun g \<in> extensional (Dom g)"
using assms 3 Fun_mapsto by simp
moreover have "Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) \<in> extensional (Dom g)"
using 5 par Fun_mapsto by (metis (no_types, lifting) Int_iff)
ultimately show "x \<notin> set (prod a b) \<Longrightarrow>
Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) x = Fun g x"
using 4 extensional_arb [of "Fun g" "Dom g" x]
extensional_arb [of "Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0'))" "Dom g" x]
by force
qed
moreover have "x \<in> set (prod a b) \<Longrightarrow>
Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) x = Fun g x"
proof -
assume x: "x \<in> set (prod a b)"
have 6: "Dom (tuple \<Lambda>_pr1' b_pr0') = set (prod a b)"
using assms 4 tuple' par
by (metis (no_types, lifting) in_homE)
have "Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) x =
Fun (eval b c) (Fun (tuple \<Lambda>_pr1' b_pr0') x)"
proof -
have "Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) x =
(Fun (eval b c) \<circ> Fun (tuple \<Lambda>_pr1' b_pr0')) x"
using assms par x 6 Fun_comp [of "eval b c" "tuple \<Lambda>_pr1' b_pr0'"] by auto
also have "... = Fun (eval b c) (Fun (tuple \<Lambda>_pr1' b_pr0') x)"
by simp
finally show ?thesis by blast
qed
also have "... = restrict (\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
(set (HF'.prod (exp b c) b))
(Fun (tuple \<Lambda>_pr1' b_pr0') x)"
using assms Fun_eval by simp
also have "... = (\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
(Fun (tuple \<Lambda>_pr1' b_pr0') x)"
proof -
have "Fun (tuple \<Lambda>_pr1' b_pr0') x \<in> set (HF'.prod (exp b c) b)"
proof -
have "x \<in> Dom (tuple \<Lambda>_pr1' b_pr0')"
using x 6 by blast
moreover have "Cod (tuple \<Lambda>_pr1' b_pr0') = set (HF'.prod (exp b c) b)"
by (metis (no_types, lifting) in_homE assms(2-3) ide_exp
prod_ide_eq tuple')
moreover have "arr (tuple \<Lambda>_pr1' b_pr0')"
using tuple' by blast
ultimately show ?thesis
using tuple' Fun_mapsto [of "tuple \<Lambda>_pr1' b_pr0'"] by auto
qed
thus ?thesis
using restrict_apply by simp
qed
also have "... = (\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
(arr_of \<langle>elem_of (Fun \<Lambda>_pr1' x), elem_of (Fun b_pr0' x)\<rangle>)"
proof -
have 7: "Dom \<Lambda>_pr1' = set (prod a b)"
using assms
by (metis (no_types, lifting) 1 comp_ide_arr ideD(2)
b_pr0_def lam_pr1_eq lam_simps(2) pr0_simps(2))
moreover have "span \<Lambda>_pr1' b_pr0'"
using assms 1 b_pr0_eq lam_pr1_eq by auto
moreover have "x \<in> Dom \<Lambda>_pr1'"
using x 7 by simp
ultimately have "Fun (tuple \<Lambda>_pr1' b_pr0') x =
arr_of \<langle>elem_of (Fun \<Lambda>_pr1' x), elem_of (Fun b_pr0' x)\<rangle>"
using assms x restrict_apply Fun_tuple by simp
thus ?thesis by simp
qed
also have "... = arr_of (happ (elem_of (Fun \<Lambda>_pr1' x)) (elem_of (Fun b_pr0' x)))"
using assms by simp
also have "... = arr_of (happ (elem_of (arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b)
(ide_to_hf c) (arr_to_hfun g))
(hfst (elem_of x)))))
(elem_of (arr_of (hsnd (elem_of x)))))"
proof -
have "Fun b_pr0' x = arr_of (hsnd (elem_of x))"
using assms x Fun_pr0 b_pr0'_def by simp
moreover have "Fun \<Lambda>_pr1' x =
arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c)
(arr_to_hfun g))
(hfst (elem_of x)))"
proof -
have "Fun \<Lambda>_pr1' x =
restrict (Fun (\<Lambda> a b c g) o Fun (pr1 a b)) (Dom (pr1 a b)) x"
using assms x Fun_pr1 Fun_comp lam_pr1_eq Fun_lam_pr1 pr1_simps(1-2)
by presburger
also have "... = Fun (\<Lambda> a b c g) (Fun (pr1 a b) x)"
using assms x restrict_apply Fun_lam_pr1 Fun_pr1 calculation lam_pr1_eq
by auto
also have "... = restrict (\<lambda>x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b)
(ide_to_hf c) (arr_to_hfun g))
(elem_of x)))
(set a)
(Fun (pr1 a b) x)"
using assms x Fun_lam by simp
also have "... = arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c)
(arr_to_hfun g))
(elem_of (Fun (pr1 a b) x)))"
proof -
have "Fun (pr1 a b) x \<in> set a"
proof -
have "x \<in> Dom (pr1 a b)"
using assms x pr1_simps(1-2) by auto
moreover have "Cod (pr1 a b) = set a"
using assms HF'.cod_pr1 pr1_simps(1) by auto
moreover have "arr (pr1 a b)"
using assms arr_char by blast
ultimately show ?thesis
using Fun_mapsto [of "pr1 a b"] by auto
qed
thus ?thesis
using restrict_apply by simp
qed
also have "... = arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c)
(arr_to_hfun g))
(hfst (elem_of x)))"
using assms x Fun_pr1 Fun_lam [of a b c g] by simp
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
also have "... = arr_of (happ (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c)
(arr_to_hfun g))
(hfst (elem_of x)))
(hsnd (elem_of x)))"
by simp
also have "... = arr_of (happ (arr_to_hfun g) (elem_of x))"
using assms x happ_hlam
by (metis (no_types, lifting) prod_def elem_of_membI HCollect_iff ide_dom
in_homE UNIV_I arr_to_hfun_in_hexp hexp_def hfst_conv hsnd_conv
ide_to_hf_hf_to_ide timesE)
also have "... = Fun g x"
using assms x 3 4 Fun_char [of g] restrict_apply [of "Fun g" "Dom g" x]
by simp
finally show ?thesis by simp
qed
ultimately show "Fun (comp (eval b c) (tuple \<Lambda>_pr1' b_pr0')) x = Fun g x"
by auto
qed
qed
finally show ?thesis by simp
qed
lemma lam_eval_prod:
assumes "ide a" and "ide b" and "ide c"
and "in_hom h a (exp b c)"
shows "\<Lambda> a b c (comp (eval b c) (HF'.prod h b)) = h"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 0: "in_hom (comp (eval b c) (HF'.prod h b)) (prod a b) c"
proof
show "in_hom (HF'.prod h b) (prod a b) (HF'.prod (exp b c) b)"
proof
show 1: "arr (HF'.prod h b)"
using assms HF'.prod_in_hom'
by (metis (no_types, lifting) ideD(1) in_homE)
show "dom (HF'.prod h b) = prod a b"
using assms 1
by (metis (no_types, lifting) HF'.prod_simps(2) ideD(1-2) in_homE prod_ide_eq)
show "cod (HF'.prod h b) = HF'.prod (exp b c) b"
using assms 1
by (metis (no_types, lifting) HF'.prod_simps(3) ideD(1,3) in_homE)
qed
show "in_hom (eval b c) (HF'.prod (exp b c) b) c"
using assms by blast
qed
have 1: "in_hom (\<Lambda> a b c (comp (eval b c) (HF'.prod h b))) a (exp b c)"
using assms 0 by blast
have 2: "Fun (comp (eval b c) (HF'.prod h b)) =
restrict (Fun (eval b c) \<circ> Fun (HF'.prod h b))
(set (HF'.prod a b))"
proof -
have "seq (eval b c) (HF'.prod h b)"
using assms 1
by (metis (no_types, lifting) 0 in_homE)
moreover have "Dom (HF'.prod h b) = set (HF'.prod a b)"
using assms
by (metis (no_types, lifting) HF'.prod_simps(2) ideD(1-2) in_homE)
ultimately show ?thesis
using assms Fun_comp [of "eval b c" "HF'.prod h b"] by simp
qed
show par: "par (\<Lambda> a b c (comp (eval b c) (HF'.prod h b))) h"
using assms 1
by (metis (no_types, lifting) in_homE)
show "Fun (\<Lambda> a b c (comp (eval b c) (HF'.prod h b))) = Fun h"
proof
fix x
show "Fun (\<Lambda> a b c (comp (eval b c) (HF'.prod h b))) x = Fun h x"
proof -
have "x \<notin> set a \<Longrightarrow> ?thesis"
using assms 1 Fun_mapsto
extensional_arb [of "Fun h" "set a" x]
extensional_arb [of "Fun (\<Lambda> a b c (comp (eval b c) (HF'.prod h b)))"
"set a" x]
by (metis (no_types, lifting) 0 Int_iff lam_simps(2) par)
moreover have "x \<in> set a \<Longrightarrow> ?thesis"
proof -
assume x: "x \<in> set a"
have 3: "dom (comp (eval b c) (HF'.prod h b)) = HF'.prod a b"
using assms 0 in_homE prod_ide_eq by auto
have 4: "cod (comp (eval b c) (HF'.prod h b)) = c"
using assms 0 by blast
have 5: "dom (comp (eval b c) (HF'.prod h b)) = HF'.prod a b"
using assms 3
by (metis (mono_tags, lifting))
have 6: "cod (comp (eval b c) (HF'.prod h b)) = c"
using assms 4 by (metis (no_types, lifting))
have 7: "arr_to_hfun (comp (eval b c) (HF'.prod h b)) =
\<lbrace>xy \<^bold>\<in> ide_to_hf (HF'.prod a b) * ide_to_hf c.
hsnd xy = elem_of (Fun (comp (eval b c) (HF'.prod h b)) (arr_of (hfst xy)))\<rbrace>"
unfolding arr_to_hfun_def
using 2 5 6 by metis
have "Fun (\<Lambda> a b c (comp (eval b c) (HF'.prod h b))) x =
arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c)
(arr_to_hfun (comp (eval b c) (HF'.prod h b))))
(elem_of x))"
using assms 0 x Fun_lam by auto
also have "... = arr_of \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle>
\<^bold>\<in> arr_to_hfun (comp (eval b c) (HF'.prod h b))\<rbrace>"
proof -
have "seq (eval b c) (HF'.prod h b)"
using assms 0 by blast
moreover have "ide_to_hf (dom (comp (eval b c) (HF'.prod h b))) =
ide_to_hf a * ide_to_hf b"
using assms 1 3
by (metis (no_types, lifting) prod_def UNIV_I ide_to_hf_hf_to_ide prod_ide_eq)
moreover have "ide_to_hf (cod (comp (eval b c) (HF'.prod h b))) = ide_to_hf c"
using assms 2 4 by auto
ultimately show ?thesis
using assms 0 x happ_hlam(3) elem_of_membI
hfun_arr_to_hfun [of "comp (eval b c) (HF'.prod h b)"]
by simp
qed
also have "... = arr_of \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))\<rbrace>"
proof -
have "\<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle>
\<^bold>\<in> arr_to_hfun (comp (eval b c) (HF'.prod h b))\<rbrace> =
\<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))\<rbrace>"
proof
fix yz
show "yz \<^bold>\<in> \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle>
\<^bold>\<in> arr_to_hfun (comp (eval b c) (HF'.prod h b))\<rbrace> \<longleftrightarrow>
yz \<^bold>\<in> \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))\<rbrace>"
proof -
have "yz \<^bold>\<in> ide_to_hf b * ide_to_hf c \<Longrightarrow>
\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle> \<^bold>\<in> arr_to_hfun (comp (eval b c) (HF'.prod h b))
\<longleftrightarrow> hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))"
proof -
assume yz: "yz \<^bold>\<in> ide_to_hf b * ide_to_hf c"
have "\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle>
\<^bold>\<in> arr_to_hfun (comp (eval b c) (HF'.prod h b))
\<longleftrightarrow>
\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle> \<^bold>\<in> ide_to_hf (HF'.prod a b) * ide_to_hf c \<and>
hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))"
using 7 by auto
moreover have "\<langle>\<langle>elem_of x, hfst yz\<rangle>, hsnd yz\<rangle>
\<^bold>\<in> ide_to_hf (prod a b) * ide_to_hf c"
proof -
have "\<langle>elem_of x, hfst yz\<rangle> \<^bold>\<in> ide_to_hf (HF'.prod a b)"
using assms x yz
by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I hfst_conv
ide_to_hf_hf_to_ide prod_ide_eq timesE times_iff)
thus ?thesis
using yz assms(1-2) prod_ide_eq by auto
qed
ultimately show ?thesis
using assms(1-2) prod_ide_eq by auto
qed
thus ?thesis by auto
qed
qed
thus ?thesis by simp
qed
also have "... = arr_of \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>"
proof -
have "\<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c.
hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))\<rbrace> =
\<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>"
proof -
have "\<And>yz. yz \<^bold>\<in> ide_to_hf b * ide_to_hf c \<Longrightarrow>
hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))
\<longleftrightarrow>
yz \<^bold>\<in> elem_of (Fun h x)"
proof -
fix yz
assume yz: "yz \<^bold>\<in> ide_to_hf b * ide_to_hf c"
have 7: "arr_of \<langle>elem_of x, hfst yz\<rangle> \<in> set (HF'.prod a b)"
using assms x yz arr_of_membI
by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I hfst_conv
ide_to_hf_hf_to_ide prod_ide_eq timesE times_iff)
have 8: "Fun h x \<in> set (exp b c)"
proof -
have "Fun h x \<in> Cod h"
using assms x Fun_mapsto by blast
moreover have "Cod h = set (exp b c)"
using assms 0 lam_simps(3) par by auto
ultimately show ?thesis by blast
qed
show "hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>))
\<longleftrightarrow>
yz \<^bold>\<in> elem_of (Fun h x)"
proof -
have "Fun (comp (eval b c) (HF'.prod h b)) (arr_of \<langle>elem_of x, hfst yz\<rangle>) =
arr_of (happ (elem_of (Fun h x)) (hfst yz))"
proof -
have "Fun (comp (eval b c) (HF'.prod h b)) (arr_of \<langle>elem_of x, hfst yz\<rangle>) =
restrict (Fun (eval b c) \<circ> Fun (HF'.prod h b))
(set (HF'.prod a b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>)"
using assms x yz 2 by simp
also have "... = Fun (eval b c)
(Fun (HF'.prod h b) (arr_of \<langle>elem_of x, hfst yz\<rangle>))"
using 7 by simp
also have "... = Fun (eval b c)
(arr_of \<langle>elem_of (Fun h x),
elem_of (Fun b (arr_of (hfst yz)))\<rangle>)"
proof -
have "Fun (HF'.prod h b) (arr_of \<langle>elem_of x, hfst yz\<rangle>) =
arr_of \<langle>elem_of (Fun h x), elem_of (Fun b (arr_of (hfst yz)))\<rangle>"
proof -
have "Fun (HF'.prod h b) (arr_of \<langle>elem_of x, hfst yz\<rangle>) =
arr_of \<langle>elem_of (Fun h (arr_of (hfst (elem_of (arr_of \<langle>elem_of x, hfst yz\<rangle>))))),
elem_of (Fun b (arr_of (hsnd (elem_of (arr_of \<langle>elem_of x, hfst yz\<rangle>)))))\<rangle>"
proof -
have "arr_of \<langle>elem_of x, hfst yz\<rangle> \<in> set (prod (dom h) (dom b))"
using assms x yz 7
by (metis (no_types, lifting) ideD(2) in_homE prod_ide_eq)
thus ?thesis
using assms x yz Fun_prod ideD(1) by blast
qed
also have "... = arr_of \<langle>elem_of (Fun h (arr_of (elem_of x))),
elem_of (Fun b (arr_of (hfst yz)))\<rangle>"
using assms x yz by simp
also have "... = arr_of \<langle>elem_of (Fun h x), elem_of (Fun b (arr_of (hfst yz)))\<rangle>"
using assms(1) set_ideD(1) x by force
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = Fun (eval b c) (arr_of \<langle>elem_of (Fun h x), hfst yz\<rangle>)"
using assms x yz Fun_ide ide_char arr_of_membI by auto
also have "... = restrict (\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
(set (HF'.prod (exp b c) b))
(arr_of \<langle>elem_of (Fun h x), hfst yz\<rangle>)"
using assms Fun_eval [of b c] by simp
also have "... = (\<lambda>x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))
(arr_of \<langle>elem_of (Fun h x), hfst yz\<rangle>)"
proof -
have "arr_of \<langle>elem_of (Fun h x), hfst yz\<rangle>
\<in> set (HF'.prod (exp b c) b)"
proof -
have 1: "ide_to_hf (HF'.prod (exp b c) b) =
HF (elem_of ` set (HF'.prod (exp b c) b))"
unfolding ide_to_hf_def by blast
have "\<langle>elem_of (Fun h x), hfst yz\<rangle>
\<^bold>\<in> HF (elem_of ` set (HF'.prod (exp b c) b))"
using assms x yz 1 8 Fun_mapsto [of h]
by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I
hfst_conv ide_exp ide_to_hf_hf_to_ide prod_ide_eq timesE times_iff)
thus ?thesis
using assms x yz 1 arr_of_membI [of "\<langle>elem_of (Fun h x), hfst yz\<rangle>"]
by auto
qed
thus ?thesis by simp
qed
also have "... = arr_of (happ (elem_of (Fun h x)) (hfst yz))"
by simp
finally show ?thesis by simp
qed
hence 9: "elem_of (Fun (comp (eval b c) (HF'.prod h b))
(arr_of \<langle>elem_of x, hfst yz\<rangle>)) =
happ (elem_of (Fun h x)) (hfst yz)"
by simp
show ?thesis
proof -
have "hsnd yz = happ (elem_of (Fun h x)) (hfst yz)
\<longleftrightarrow> yz \<^bold>\<in> elem_of (Fun h x)"
proof
have 10: "\<exists>!z. \<langle>hfst yz, z\<rangle> \<^bold>\<in> elem_of (Fun h x)"
proof -
have "hfun (ide_to_hf b) (ide_to_hf c) (elem_of (Fun h x))"
using assms x 8
by (metis (no_types, lifting) elem_of_membI HCollect_iff UNIV_I
exp_def hexp_def ide_exp ide_to_hf_hf_to_ide)
thus ?thesis
using assms yz
hfunE [of "ide_to_hf b" "ide_to_hf c" "elem_of (Fun h x)"]
by (metis (no_types, lifting) hfst_conv timesE)
qed
show "yz \<^bold>\<in> elem_of (Fun h x)
\<Longrightarrow> hsnd yz = happ (elem_of (Fun h x)) (hfst yz)"
proof -
assume yz1: "yz \<^bold>\<in> elem_of (Fun h x)"
show "hsnd yz = happ (elem_of (Fun h x)) (hfst yz)"
unfolding app_def
using assms x yz yz1 10 hfun_arr_to_hfun arr_to_hfun_def
the1_equality
[of "\<lambda>y. \<langle>hfst yz, y\<rangle> \<^bold>\<in> elem_of (Fun h x)" "hsnd yz"]
by (metis (no_types, lifting) hfst_conv hsnd_conv timesE)
qed
show "hsnd yz = happ (elem_of (Fun h x)) (hfst yz)
\<Longrightarrow> yz \<^bold>\<in> elem_of (Fun h x)"
unfolding app_def
using assms x yz 10
theI' [of "\<lambda>y. \<langle>hfst yz, y\<rangle> \<^bold>\<in> elem_of (Fun h x)"]
by (metis (no_types, lifting) hfst_conv hsnd_conv timesE)
qed
thus ?thesis
using 9 by simp
qed
qed
qed
thus ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = Fun h x"
proof -
have H: "Fun h x = restrict (\<lambda>x. arr_of (happ (arr_to_hfun h) (elem_of x))) (Dom h) x"
proof -
have "arr h"
using assms by blast
thus ?thesis
using assms x Fun_char by simp
qed
also have "... = arr_of (happ (arr_to_hfun h) (elem_of x))"
using assms x par
by (metis (no_types, lifting) 0 lam_simps(2) restrict_apply)
also have "... = arr_of (THE g. \<langle>elem_of x, g\<rangle> \<^bold>\<in> arr_to_hfun h)"
using app_def by simp
also have "... = arr_of \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>"
proof -
have ex_un_g: "\<exists>!g. \<langle>elem_of x, g\<rangle> \<^bold>\<in> arr_to_hfun h"
using assms x arr_to_hfun_def hfun_arr_to_hfun
hfunE [of "ide_to_hf a" "ide_to_hf (exp b c)" "arr_to_hfun h"]
by (metis (no_types, lifting) elem_of_membI in_homE)
moreover have
"\<langle>elem_of x, \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>\<rangle>
\<^bold>\<in> arr_to_hfun h"
proof -
have "elem_of (Fun h x) =
\<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>"
proof
fix yz
show "yz \<^bold>\<in> elem_of (Fun h x) \<longleftrightarrow>
yz \<^bold>\<in> \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>"
proof
show "yz \<^bold>\<in> elem_of (Fun h x)
\<Longrightarrow> yz \<^bold>\<in> \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>"
proof -
assume yz: "yz \<^bold>\<in> elem_of (Fun h x)"
have "yz \<^bold>\<in> ide_to_hf b * ide_to_hf c"
proof -
have "elem_of (Fun h x) \<^bold>\<in> hexp (ide_to_hf b) (ide_to_hf c)"
proof -
have "ide (hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c)))"
using assms exp_def ide_exp by auto
moreover have
"Fun h x \<in> set (hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c)))"
proof -
have "Fun h x \<in> Cod h"
using assms x Fun_mapsto by blast
moreover have
"Cod h = set (hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c)))"
using assms 0 exp_def lam_simps(3) par by auto
ultimately show ?thesis by blast
qed
ultimately show ?thesis
using elem_of_membI [of "hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c))"
"Fun h x"]
by (simp add: ide_to_hf_hf_to_ide)
qed
thus ?thesis
using assms yz hexp_def by auto
qed
thus ?thesis
using assms x yz by blast
qed
show "yz \<^bold>\<in> \<lbrace>yz \<^bold>\<in> ide_to_hf b * ide_to_hf c. yz \<^bold>\<in> elem_of (Fun h x)\<rbrace>
\<Longrightarrow> yz \<^bold>\<in> elem_of (Fun h x)"
using assms by simp
qed
qed
moreover have "arr_of (elem_of x) = x"
using arr_of_elem_of assms(1) set_ideD(1) x by blast
ultimately show ?thesis
using assms x arr_to_hfun_def ex_un_g by auto
qed
ultimately show ?thesis
using assms x theI' [of "\<lambda>g. \<langle>elem_of x, g\<rangle> \<^bold>\<in> arr_to_hfun h"]
by fastforce
qed
finally show ?thesis
using assms x by simp
qed
finally show ?thesis by simp
qed
ultimately show "Fun (\<Lambda> a b c (comp (eval b c) (HF'.prod h b))) x = Fun h x"
by blast
qed
qed
qed
section "The Main Results"
interpretation cartesian_closed_category comp
proof -
interpret elementary_cartesian_closed_category comp pr0 pr1
some_terminal trm exp eval \<Lambda>
using ide_exp eval_in_hom lam_in_hom prod_ide_eq eval_prod_lam lam_eval_prod
by unfold_locales auto
show "cartesian_closed_category comp"
using is_cartesian_closed_category by simp
qed
theorem is_cartesian_closed_category:
shows "cartesian_closed_category comp"
..
theorem is_category_with_finite_limits:
shows "category_with_finite_limits comp"
proof
fix J :: "'j comp"
assume J: "category J"
interpret J: category J
using J by simp
assume finite: "finite (Collect J.arr)"
have "has_products (Collect J.ide)"
proof -
have "Collect J.ide \<noteq> UNIV"
using J.not_arr_null by blast
moreover have "finite (Collect J.ide)"
proof -
have "Collect J.ide \<subseteq> Collect J.arr"
by auto
thus ?thesis
using finite J.ideD(1) finite_subset by blast
qed
ultimately show ?thesis
using finite has_finite_products' by simp
qed
moreover have "has_products (Collect J.arr)"
proof -
have "Collect J.arr \<noteq> UNIV"
using J.not_arr_null by blast
thus ?thesis
using finite has_finite_products' by simp
qed
ultimately show "has_limits_of_shape J"
using J.category_axioms has_limits_if_has_products [of J] by simp
qed
end
end
diff --git a/thys/Category3/Limit.thy b/thys/Category3/Limit.thy
--- a/thys/Category3/Limit.thy
+++ b/thys/Category3/Limit.thy
@@ -1,6352 +1,6352 @@
(* Title: Limit
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2016
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter Limit
theory Limit
imports FreeCategory DiscreteCategory Adjunction
begin
text\<open>
This theory defines the notion of limit in terms of diagrams and cones and relates
it to the concept of a representation of a functor. The diagonal functor associated
with a diagram shape @{term J} is defined and it is shown that a right adjoint to
the diagonal functor gives limits of shape @{term J} and that a category has limits
of shape @{term J} if and only if the diagonal functor is a left adjoint functor.
Products and equalizers are defined as special cases of limits, and it is shown
that a category with equalizers has limits of shape @{term J} if it has products
indexed by the sets of objects and arrows of @{term J}.
The existence of limits in a set category is investigated, and it is shown that
every set category has equalizers and that a set category @{term S} has @{term I}-indexed
products if and only if the universe of @{term S} ``admits @{term I}-indexed tupling.''
The existence of limits in functor categories is also developed, showing that
limits in functor categories are ``determined pointwise'' and that a functor category
@{term "[A, B]"} has limits of shape @{term J} if @{term B} does.
Finally, it is shown that the Yoneda functor preserves limits.
This theory concerns itself only with limits; I have made no attempt to consider colimits.
Although it would be possible to rework the entire development in dual form,
it is possible that there is a more efficient way to dualize at least parts of it without
repeating all the work. This is something that deserves further thought.
\<close>
section "Representations of Functors"
text\<open>
A representation of a contravariant functor \<open>F: Cop \<rightarrow> S\<close>, where @{term S}
is a set category that is the target of a hom-functor for @{term C}, consists of
an object @{term a} of @{term C} and a natural isomorphism @{term "\<Phi>: Y a \<rightarrow> F"},
where \<open>Y: C \<rightarrow> [Cop, S]\<close> is the Yoneda functor.
\<close>
locale representation_of_functor =
C: category C +
Cop: dual_category C +
S: set_category S setp +
F: "functor" Cop.comp S F +
Hom: hom_functor C S setp \<phi> +
Ya: yoneda_functor_fixed_object C S setp \<phi> a +
natural_isomorphism Cop.comp S \<open>Ya.Y a\<close> F \<Phi>
for C :: "'c comp" (infixr "\<cdot>" 55)
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and setp :: "'s set \<Rightarrow> bool"
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and F :: "'c \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
begin
abbreviation Y where "Y \<equiv> Ya.Y"
abbreviation \<psi> where "\<psi> \<equiv> Hom.\<psi>"
end
text\<open>
Two representations of the same functor are uniquely isomorphic.
\<close>
locale two_representations_one_functor =
C: category C +
Cop: dual_category C +
S: set_category S setp +
F: set_valued_functor Cop.comp S setp F +
yoneda_functor C S setp \<phi> +
Ya: yoneda_functor_fixed_object C S setp \<phi> a +
Ya': yoneda_functor_fixed_object C S setp \<phi> a' +
\<Phi>: representation_of_functor C S setp \<phi> F a \<Phi> +
\<Phi>': representation_of_functor C S setp \<phi> F a' \<Phi>'
for C :: "'c comp" (infixr "\<cdot>" 55)
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and setp :: "'s set \<Rightarrow> bool"
and F :: "'c \<Rightarrow> 's"
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
and a' :: 'c
and \<Phi>' :: "'c \<Rightarrow> 's"
begin
interpretation \<Psi>: inverse_transformation Cop.comp S \<open>Y a\<close> F \<Phi> ..
interpretation \<Psi>': inverse_transformation Cop.comp S \<open>Y a'\<close> F \<Phi>' ..
interpretation \<Phi>\<Psi>': vertical_composite Cop.comp S \<open>Y a\<close> F \<open>Y a'\<close> \<Phi> \<Psi>'.map ..
interpretation \<Phi>'\<Psi>: vertical_composite Cop.comp S \<open>Y a'\<close> F \<open>Y a\<close> \<Phi>' \<Psi>.map ..
lemma are_uniquely_isomorphic:
shows "\<exists>!\<phi>. \<guillemotleft>\<phi> : a \<rightarrow> a'\<guillemotright> \<and> C.iso \<phi> \<and> map \<phi> = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
proof -
interpret \<Phi>\<Psi>': natural_isomorphism Cop.comp S \<open>Y a\<close> \<open>Y a'\<close> \<Phi>\<Psi>'.map
using \<Phi>.natural_isomorphism_axioms \<Psi>'.natural_isomorphism_axioms
natural_isomorphisms_compose
by blast
interpret \<Phi>'\<Psi>: natural_isomorphism Cop.comp S \<open>Y a'\<close> \<open>Y a\<close> \<Phi>'\<Psi>.map
using \<Phi>'.natural_isomorphism_axioms \<Psi>.natural_isomorphism_axioms
natural_isomorphisms_compose
by blast
interpret \<Phi>\<Psi>'_\<Phi>'\<Psi>: inverse_transformations Cop.comp S \<open>Y a\<close> \<open>Y a'\<close> \<Phi>\<Psi>'.map \<Phi>'\<Psi>.map
proof
fix x
assume X: "Cop.ide x"
show "S.inverse_arrows (\<Phi>\<Psi>'.map x) (\<Phi>'\<Psi>.map x)"
using S.inverse_arrows_compose S.inverse_arrows_sym X \<Phi>'\<Psi>.map_simp_ide
\<Phi>\<Psi>'.map_simp_ide \<Psi>'.inverts_components \<Psi>.inverts_components
by force
qed
have "Cop_S.inverse_arrows (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)
(Cop_S.MkArr (Y a') (Y a) \<Phi>'\<Psi>.map)"
proof -
have Ya: "functor Cop.comp S (Y a)" ..
have Ya': "functor Cop.comp S (Y a')" ..
have \<Phi>\<Psi>': "natural_transformation Cop.comp S (Y a) (Y a') \<Phi>\<Psi>'.map" ..
have \<Phi>'\<Psi>: "natural_transformation Cop.comp S (Y a') (Y a) \<Phi>'\<Psi>.map" ..
show ?thesis
by (metis (no_types, lifting) Cop_S.arr_MkArr Cop_S.comp_MkArr Cop_S.ide_MkIde
Cop_S.inverse_arrows_def Ya'.functor_axioms Ya.functor_axioms
\<Phi>'\<Psi>.natural_transformation_axioms \<Phi>\<Psi>'.natural_transformation_axioms
\<Phi>\<Psi>'_\<Phi>'\<Psi>.inverse_transformations_axioms inverse_transformations_inverse(1-2)
mem_Collect_eq)
qed
hence 3: "Cop_S.iso (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)" using Cop_S.isoI by blast
hence "\<exists>f. \<guillemotleft>f : a \<rightarrow> a'\<guillemotright> \<and> map f = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
using Ya.ide_a Ya'.ide_a is_full Y_def Cop_S.iso_is_arr full_functor.is_full
Cop_S.MkArr_in_hom \<Phi>\<Psi>'.natural_transformation_axioms preserves_ide
by force
from this obtain \<phi>
where \<phi>: "\<guillemotleft>\<phi> : a \<rightarrow> a'\<guillemotright> \<and> map \<phi> = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
by blast
show ?thesis
by (metis 3 C.in_homE \<phi> is_faithful reflects_iso)
qed
end
section "Diagrams and Cones"
text\<open>
A \emph{diagram} in a category @{term C} is a functor \<open>D: J \<rightarrow> C\<close>.
We refer to the category @{term J} as the diagram \emph{shape}.
Note that in the usual expositions of category theory that use set theory
as their foundations, the shape @{term J} of a diagram is required to be
a ``small'' category, where smallness means that the collection of objects
of @{term J}, as well as each of the ``homs,'' is a set.
However, in HOL there is no class of all sets, so it is not meaningful
to speak of @{term J} as ``small'' in any kind of absolute sense.
There is likely a meaningful notion of smallness of @{term J}
\emph{relative to} @{term C} (the result below that states that a set
category has @{term I}-indexed products if and only if its universe
``admits @{term I}-indexed tuples'' is suggestive of how this might
be defined), but I haven't fully explored this idea at present.
\<close>
locale diagram =
C: category C +
J: category J +
"functor" J C D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
begin
notation J.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J _\<guillemotright>")
end
lemma comp_diagram_functor:
assumes "diagram J C D" and "functor J' J F"
shows "diagram J' C (D o F)"
by (meson assms(1) assms(2) diagram_def functor.axioms(1) functor_comp)
text\<open>
A \emph{cone} over a diagram \<open>D: J \<rightarrow> C\<close> is a natural transformation
from a constant functor to @{term D}. The value of the constant functor is
the \emph{apex} of the cone.
\<close>
locale cone =
C: category C +
J: category J +
D: diagram J C D +
A: constant_functor J C a +
natural_transformation J C A.map D \<chi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and a :: 'c
and \<chi> :: "'j \<Rightarrow> 'c"
begin
lemma ide_apex:
shows "C.ide a"
using A.value_is_ide by auto
lemma component_in_hom:
assumes "J.arr j"
shows "\<guillemotleft>\<chi> j : a \<rightarrow> D (J.cod j)\<guillemotright>"
using assms by auto
end
text\<open>
A cone over diagram @{term D} is transformed into a cone over diagram @{term "D o F"}
by pre-composing with @{term F}.
\<close>
lemma comp_cone_functor:
assumes "cone J C D a \<chi>" and "functor J' J F"
shows "cone J' C (D o F) a (\<chi> o F)"
proof -
interpret \<chi>: cone J C D a \<chi> using assms(1) by auto
interpret F: "functor" J' J F using assms(2) by auto
interpret A': constant_functor J' C a
using \<chi>.A.value_is_ide by unfold_locales auto
have 1: "\<chi>.A.map o F = A'.map"
using \<chi>.A.map_def A'.map_def \<chi>.J.not_arr_null by auto
interpret \<chi>': natural_transformation J' C A'.map \<open>D o F\<close> \<open>\<chi> o F\<close>
using 1 horizontal_composite F.as_nat_trans.natural_transformation_axioms
\<chi>.natural_transformation_axioms
by fastforce
show "cone J' C (D o F) a (\<chi> o F)" ..
qed
text\<open>
A cone over diagram @{term D} can be transformed into a cone over a diagram @{term D'}
by post-composing with a natural transformation from @{term D} to @{term D'}.
\<close>
lemma vcomp_transformation_cone:
assumes "cone J C D a \<chi>"
and "natural_transformation J C D D' \<tau>"
shows "cone J C D' a (vertical_composite.map J C \<chi> \<tau>)"
by (meson assms cone.axioms(4-5) cone.intro diagram.intro natural_transformation.axioms(1-4)
vertical_composite.intro vertical_composite.is_natural_transformation)
context "functor"
begin
lemma preserves_diagrams:
fixes J :: "'j comp"
assumes "diagram J A D"
shows "diagram J B (F o D)"
by (meson assms diagram_def functor_axioms functor_comp functor_def)
lemma preserves_cones:
fixes J :: "'j comp"
assumes "cone J A D a \<chi>"
shows "cone J B (F o D) (F a) (F o \<chi>)"
proof -
interpret \<chi>: cone J A D a \<chi> using assms by auto
interpret Fa: constant_functor J B \<open>F a\<close>
using \<chi>.ide_apex by unfold_locales auto
have 1: "F o \<chi>.A.map = Fa.map"
proof
fix f
show "(F \<circ> \<chi>.A.map) f = Fa.map f"
using is_extensional Fa.is_extensional \<chi>.A.is_extensional
by (cases "\<chi>.J.arr f", simp_all)
qed
interpret \<chi>': natural_transformation J B Fa.map \<open>F o D\<close> \<open>F o \<chi>\<close>
using 1 horizontal_composite \<chi>.natural_transformation_axioms
as_nat_trans.natural_transformation_axioms
by fastforce
show "cone J B (F o D) (F a) (F o \<chi>)" ..
qed
end
context diagram
begin
abbreviation cone
where "cone a \<chi> \<equiv> Limit.cone J C D a \<chi>"
abbreviation cones :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) set"
where "cones a \<equiv> { \<chi>. cone a \<chi> }"
text\<open>
An arrow @{term "f \<in> C.hom a' a"} induces by composition a transformation from
cones with apex @{term a} to cones with apex @{term a'}. This transformation
is functorial in @{term f}.
\<close>
abbreviation cones_map :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) \<Rightarrow> ('j \<Rightarrow> 'c)"
where "cones_map f \<equiv> (\<lambda>\<chi> \<in> cones (C.cod f). \<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null)"
lemma cones_map_mapsto:
assumes "C.arr f"
shows "cones_map f \<in>
extensional (cones (C.cod f)) \<inter> (cones (C.cod f) \<rightarrow> cones (C.dom f))"
proof
show "cones_map f \<in> extensional (cones (C.cod f))" by blast
show "cones_map f \<in> cones (C.cod f) \<rightarrow> cones (C.dom f)"
proof
fix \<chi>
assume "\<chi> \<in> cones (C.cod f)"
hence \<chi>: "cone (C.cod f) \<chi>" by auto
interpret \<chi>: cone J C D \<open>C.cod f\<close> \<chi> using \<chi> by auto
interpret B: constant_functor J C \<open>C.dom f\<close>
using assms by unfold_locales auto
have "cone (C.dom f) (\<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null)"
using assms B.value_is_ide \<chi>.is_natural_1 \<chi>.is_natural_2
apply (unfold_locales, auto)
using \<chi>.is_natural_1
apply (metis C.comp_assoc)
using \<chi>.is_natural_2 C.comp_arr_dom
by (metis J.arr_cod_iff_arr J.cod_cod C.comp_assoc)
thus "(\<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null) \<in> cones (C.dom f)" by auto
qed
qed
lemma cones_map_ide:
assumes "\<chi> \<in> cones a"
shows "cones_map a \<chi> = \<chi>"
proof -
interpret \<chi>: cone J C D a \<chi> using assms by auto
show ?thesis
proof
fix j
show "cones_map a \<chi> j = \<chi> j"
using assms \<chi>.A.value_is_ide \<chi>.preserves_hom C.comp_arr_dom \<chi>.is_extensional
by (cases "J.arr j", auto)
qed
qed
lemma cones_map_comp:
assumes "C.seq f g"
shows "cones_map (f \<cdot> g) = restrict (cones_map g o cones_map f) (cones (C.cod f))"
proof (intro restr_eqI)
show "cones (C.cod (f \<cdot> g)) = cones (C.cod f)" using assms by simp
show "\<And>\<chi>. \<chi> \<in> cones (C.cod (f \<cdot> g)) \<Longrightarrow>
(\<lambda>j. if J.arr j then \<chi> j \<cdot> f \<cdot> g else C.null) = (cones_map g o cones_map f) \<chi>"
proof -
fix \<chi>
assume \<chi>: "\<chi> \<in> cones (C.cod (f \<cdot> g))"
show "(\<lambda>j. if J.arr j then \<chi> j \<cdot> f \<cdot> g else C.null) = (cones_map g o cones_map f) \<chi>"
proof -
have "((cones_map g) o (cones_map f)) \<chi> = cones_map g (cones_map f \<chi>)"
by force
also have "... = (\<lambda>j. if J.arr j then
(\<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null) j \<cdot> g else C.null)"
proof
fix j
have "cone (C.dom f) (cones_map f \<chi>)"
using assms \<chi> cones_map_mapsto by (elim C.seqE, force)
thus "cones_map g (cones_map f \<chi>) j =
(if J.arr j then C (if J.arr j then \<chi> j \<cdot> f else C.null) g else C.null)"
using \<chi> assms by auto
qed
also have "... = (\<lambda>j. if J.arr j then \<chi> j \<cdot> f \<cdot> g else C.null)"
using C.comp_assoc by fastforce
finally show ?thesis by auto
qed
qed
qed
end
text\<open>
Changing the apex of a cone by pre-composing with an arrow @{term f} commutes
with changing the diagram of a cone by post-composing with a natural transformation.
\<close>
lemma cones_map_vcomp:
assumes "diagram J C D" and "diagram J C D'"
and "natural_transformation J C D D' \<tau>"
and "cone J C D a \<chi>"
and f: "partial_magma.in_hom C f a' a"
shows "diagram.cones_map J C D' f (vertical_composite.map J C \<chi> \<tau>)
= vertical_composite.map J C (diagram.cones_map J C D f \<chi>) \<tau>"
proof -
interpret D: diagram J C D using assms(1) by auto
interpret D': diagram J C D' using assms(2) by auto
interpret \<tau>: natural_transformation J C D D' \<tau> using assms(3) by auto
interpret \<chi>: cone J C D a \<chi> using assms(4) by auto
interpret \<tau>o\<chi>: vertical_composite J C \<chi>.A.map D D' \<chi> \<tau> ..
interpret \<tau>o\<chi>: cone J C D' a \<tau>o\<chi>.map ..
interpret \<chi>f: cone J C D a' \<open>D.cones_map f \<chi>\<close>
using f \<chi>.cone_axioms D.cones_map_mapsto by blast
interpret \<tau>o\<chi>f: vertical_composite J C \<chi>f.A.map D D' \<open>D.cones_map f \<chi>\<close> \<tau> ..
interpret \<tau>o\<chi>_f: cone J C D' a' \<open>D'.cones_map f \<tau>o\<chi>.map\<close>
using f \<tau>o\<chi>.cone_axioms D'.cones_map_mapsto [of f] by blast
write C (infixr "\<cdot>" 55)
show "D'.cones_map f \<tau>o\<chi>.map = \<tau>o\<chi>f.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J C \<chi>f.A.map D' (D'.cones_map f \<tau>o\<chi>.map)" ..
show "natural_transformation J C \<chi>f.A.map D' \<tau>o\<chi>f.map" ..
show "\<And>j. D.J.ide j \<Longrightarrow> D'.cones_map f \<tau>o\<chi>.map j = \<tau>o\<chi>f.map j"
proof -
fix j
assume j: "D.J.ide j"
have "D'.cones_map f \<tau>o\<chi>.map j = \<tau>o\<chi>.map j \<cdot> f"
using f \<tau>o\<chi>.cone_axioms \<tau>o\<chi>.map_simp_2 \<tau>o\<chi>.is_extensional by auto
also have "... = (\<tau> j \<cdot> \<chi> (D.J.dom j)) \<cdot> f"
using j \<tau>o\<chi>.map_simp_2 by simp
also have "... = \<tau> j \<cdot> \<chi> (D.J.dom j) \<cdot> f"
using D.C.comp_assoc by simp
also have "... = \<tau>o\<chi>f.map j"
using j f \<chi>.cone_axioms \<tau>o\<chi>f.map_simp_2 by auto
finally show "D'.cones_map f \<tau>o\<chi>.map j = \<tau>o\<chi>f.map j" by auto
qed
qed
qed
text\<open>
Given a diagram @{term D}, we can construct a contravariant set-valued functor,
which takes each object @{term a} of @{term C} to the set of cones over @{term D}
with apex @{term a}, and takes each arrow @{term f} of @{term C} to the function
on cones over @{term D} induced by pre-composition with @{term f}.
For this, we need to introduce a set category @{term S} whose universe is large
enough to contain all the cones over @{term D}, and we need to have an explicit
correspondence between cones and elements of the universe of @{term S}.
A replete set category @{term S} equipped with an injective mapping
@{term_type "\<iota> :: ('j => 'c) => 's"} serves this purpose.
\<close>
locale cones_functor =
C: category C +
Cop: dual_category C +
J: category J +
D: diagram J C D +
S: replete_concrete_set_category S UNIV \<iota>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<iota> :: "('j \<Rightarrow> 'c) \<Rightarrow> 's"
begin
notation S.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
abbreviation \<o> where "\<o> \<equiv> S.DN"
definition map :: "'c \<Rightarrow> 's"
where "map = (\<lambda>f. if C.arr f then
S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom f))
(\<iota> o D.cones_map f o \<o>)
else S.null)"
lemma map_simp [simp]:
assumes "C.arr f"
shows "map f = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom f))
(\<iota> o D.cones_map f o \<o>)"
using assms map_def by auto
lemma arr_map:
assumes "C.arr f"
shows "S.arr (map f)"
proof -
have "\<iota> o D.cones_map f o \<o> \<in> \<iota> ` D.cones (C.cod f) \<rightarrow> \<iota> ` D.cones (C.dom f)"
using assms D.cones_map_mapsto by force
thus ?thesis using assms S.UP_mapsto by auto
qed
lemma map_ide:
assumes "C.ide a"
shows "map a = S.mkIde (\<iota> ` D.cones a)"
proof -
have "map a = S.mkArr (\<iota> ` D.cones a) (\<iota> ` D.cones a) (\<iota> o D.cones_map a o \<o>)"
using assms map_simp by force
also have "... = S.mkArr (\<iota> ` D.cones a) (\<iota> ` D.cones a) (\<lambda>x. x)"
using S.UP_mapsto D.cones_map_ide by force
also have "... = S.mkIde (\<iota> ` D.cones a)"
using assms S.mkIde_as_mkArr S.UP_mapsto by blast
finally show ?thesis by auto
qed
lemma map_preserves_dom:
assumes "Cop.arr f"
shows "map (Cop.dom f) = S.dom (map f)"
using assms arr_map map_ide by auto
lemma map_preserves_cod:
assumes "Cop.arr f"
shows "map (Cop.cod f) = S.cod (map f)"
using assms arr_map map_ide by auto
lemma map_preserves_comp:
assumes "Cop.seq g f"
shows "map (g \<cdot>\<^sup>o\<^sup>p f) = map g \<cdot>\<^sub>S map f"
proof -
have "map (g \<cdot>\<^sup>o\<^sup>p f) = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom g))
((\<iota> o D.cones_map g o \<o>) o (\<iota> o D.cones_map f o \<o>))"
proof -
have 1: "S.arr (map (g \<cdot>\<^sup>o\<^sup>p f))"
using assms arr_map [of "C f g"] by simp
have "map (g \<cdot>\<^sup>o\<^sup>p f) = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom g))
(\<iota> o D.cones_map (C f g) o \<o>)"
using assms map_simp [of "C f g"] by simp
also have "... = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom g))
((\<iota> o D.cones_map g o \<o>) o (\<iota> o D.cones_map f o \<o>))"
using assms 1 calculation D.cones_map_mapsto D.cones_map_comp by auto
finally show ?thesis by blast
qed
also have "... = map g \<cdot>\<^sub>S map f"
using assms arr_map [of f] arr_map [of g] map_simp S.comp_mkArr by auto
finally show ?thesis by auto
qed
lemma is_functor:
shows "functor Cop.comp S map"
apply (unfold_locales)
using map_def arr_map map_preserves_dom map_preserves_cod map_preserves_comp
by auto
end
sublocale cones_functor \<subseteq> "functor" Cop.comp S map using is_functor by auto
sublocale cones_functor \<subseteq> set_valued_functor Cop.comp S \<open>\<lambda>A. A \<subseteq> S.Univ\<close> map ..
section Limits
subsection "Limit Cones"
text\<open>
A \emph{limit cone} for a diagram @{term D} is a cone @{term \<chi>} over @{term D}
with the universal property that any other cone @{term \<chi>'} over the diagram @{term D}
factors uniquely through @{term \<chi>}.
\<close>
locale limit_cone =
C: category C +
J: category J +
D: diagram J C D +
cone J C D a \<chi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and a :: 'c
and \<chi> :: "'j \<Rightarrow> 'c" +
assumes is_universal: "cone J C D a' \<chi>' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
begin
definition induced_arrow :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) \<Rightarrow> 'c"
where "induced_arrow a' \<chi>' = (THE f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>')"
lemma induced_arrowI:
assumes \<chi>': "\<chi>' \<in> D.cones a'"
shows "\<guillemotleft>induced_arrow a' \<chi>' : a' \<rightarrow> a\<guillemotright>"
and "D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
proof -
have "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
using assms \<chi>' is_universal by simp
hence 1: "\<guillemotleft>induced_arrow a' \<chi>' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
using theI' [of "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"] induced_arrow_def
by presburger
show "\<guillemotleft>induced_arrow a' \<chi>' : a' \<rightarrow> a\<guillemotright>" using 1 by simp
show "D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'" using 1 by simp
qed
lemma cones_map_induced_arrow:
shows "induced_arrow a' \<in> D.cones a' \<rightarrow> C.hom a' a"
and "\<And>\<chi>'. \<chi>' \<in> D.cones a' \<Longrightarrow> D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
using induced_arrowI by auto
lemma induced_arrow_cones_map:
assumes "C.ide a'"
shows "(\<lambda>f. D.cones_map f \<chi>) \<in> C.hom a' a \<rightarrow> D.cones a'"
and "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> induced_arrow a' (D.cones_map f \<chi>) = f"
proof -
have a': "C.ide a'" using assms by (simp add: cone.ide_apex)
have cone_\<chi>: "cone J C D a \<chi>" ..
show "(\<lambda>f. D.cones_map f \<chi>) \<in> C.hom a' a \<rightarrow> D.cones a'"
using cone_\<chi> D.cones_map_mapsto by blast
fix f
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>"
show "induced_arrow a' (D.cones_map f \<chi>) = f"
proof -
have "D.cones_map f \<chi> \<in> D.cones a'"
using f cone_\<chi> D.cones_map_mapsto by blast
hence "\<exists>!f'. \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f' \<chi> = D.cones_map f \<chi>"
using assms is_universal by auto
thus ?thesis
using f induced_arrow_def
the1_equality [of "\<lambda>f'. \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f' \<chi> = D.cones_map f \<chi>"]
by presburger
qed
qed
text\<open>
For a limit cone @{term \<chi>} with apex @{term a}, for each object @{term a'} the
hom-set @{term "C.hom a' a"} is in bijective correspondence with the set of cones
with apex @{term a'}.
\<close>
lemma bij_betw_hom_and_cones:
assumes "C.ide a'"
shows "bij_betw (\<lambda>f. D.cones_map f \<chi>) (C.hom a' a) (D.cones a')"
proof (intro bij_betwI)
show "(\<lambda>f. D.cones_map f \<chi>) \<in> C.hom a' a \<rightarrow> D.cones a'"
using assms induced_arrow_cones_map by blast
show "induced_arrow a' \<in> D.cones a' \<rightarrow> C.hom a' a"
using assms cones_map_induced_arrow by blast
show "\<And>f. f \<in> C.hom a' a \<Longrightarrow> induced_arrow a' (D.cones_map f \<chi>) = f"
using assms induced_arrow_cones_map by blast
show "\<And>\<chi>'. \<chi>' \<in> D.cones a' \<Longrightarrow> D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
using assms cones_map_induced_arrow by blast
qed
lemma induced_arrow_eqI:
assumes "D.cone a' \<chi>'" and "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" and "D.cones_map f \<chi> = \<chi>'"
shows "induced_arrow a' \<chi>' = f"
using assms is_universal induced_arrow_def
the1_equality [of "\<lambda>f. f \<in> C.hom a' a \<and> D.cones_map f \<chi> = \<chi>'" f]
by simp
lemma induced_arrow_self:
shows "induced_arrow a \<chi> = a"
proof -
have "\<guillemotleft>a : a \<rightarrow> a\<guillemotright> \<and> D.cones_map a \<chi> = \<chi>"
using ide_apex cone_axioms D.cones_map_ide by force
thus ?thesis using induced_arrow_eqI cone_axioms by auto
qed
end
context diagram
begin
abbreviation limit_cone
where "limit_cone a \<chi> \<equiv> Limit.limit_cone J C D a \<chi>"
text\<open>
A diagram @{term D} has object @{term a} as a limit if @{term a} is the apex
of some limit cone over @{term D}.
\<close>
abbreviation has_as_limit :: "'c \<Rightarrow> bool"
where "has_as_limit a \<equiv> (\<exists>\<chi>. limit_cone a \<chi>)"
abbreviation has_limit
where "has_limit \<equiv> (\<exists>a \<chi>. limit_cone a \<chi>)"
definition some_limit :: 'c
where "some_limit = (SOME a. \<exists>\<chi>. limit_cone a \<chi>)"
definition some_limit_cone :: "'j \<Rightarrow> 'c"
where "some_limit_cone = (SOME \<chi>. limit_cone some_limit \<chi>)"
lemma limit_cone_some_limit_cone:
assumes has_limit
shows "limit_cone some_limit some_limit_cone"
proof -
have "\<exists>a. has_as_limit a" using assms by simp
hence "has_as_limit some_limit"
using some_limit_def someI_ex [of "\<lambda>a. \<exists>\<chi>. limit_cone a \<chi>"] by simp
thus "limit_cone some_limit some_limit_cone"
using assms some_limit_cone_def someI_ex [of "\<lambda>\<chi>. limit_cone some_limit \<chi>"]
by simp
qed
lemma ex_limitE:
assumes "\<exists>a. has_as_limit a"
obtains a \<chi> where "limit_cone a \<chi>"
using assms someI_ex by blast
end
subsection "Limits by Representation"
text\<open>
A limit for a diagram D can also be given by a representation \<open>(a, \<Phi>)\<close>
of the cones functor.
\<close>
locale representation_of_cones_functor =
C: category C +
Cop: dual_category C +
J: category J +
D: diagram J C D +
S: replete_concrete_set_category S UNIV \<iota> +
Cones: cones_functor J C D S \<iota> +
Hom: hom_functor C S \<open>\<lambda>A. A \<subseteq> S.Univ\<close> \<phi> +
representation_of_functor C S S.setp \<phi> Cones.map a \<Phi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and \<iota> :: "('j \<Rightarrow> 'c) \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
subsection "Putting it all Together"
text\<open>
A ``limit situation'' combines and connects the ways of presenting a limit.
\<close>
locale limit_situation =
C: category C +
Cop: dual_category C +
J: category J +
D: diagram J C D +
S: replete_concrete_set_category S UNIV \<iota> +
Cones: cones_functor J C D S \<iota> +
Hom: hom_functor C S S.setp \<phi> +
\<Phi>: representation_of_functor C S S.setp \<phi> Cones.map a \<Phi> +
\<chi>: limit_cone J C D a \<chi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and \<iota> :: "('j \<Rightarrow> 'c) \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
and \<chi> :: "'j \<Rightarrow> 'c" +
assumes \<chi>_in_terms_of_\<Phi>: "\<chi> = S.DN (S.Fun (\<Phi> a) (\<phi> (a, a) a))"
and \<Phi>_in_terms_of_\<chi>:
"Cop.ide a' \<Longrightarrow> \<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a')
(\<lambda>x. \<iota> (D.cones_map (Hom.\<psi> (a', a) x) \<chi>))"
text (in limit_situation) \<open>
The assumption @{prop \<chi>_in_terms_of_\<Phi>} states that the universal cone @{term \<chi>} is obtained
by applying the function @{term "S.Fun (\<Phi> a)"} to the identity @{term a} of
@{term[source=true] C} (after taking into account the necessary coercions).
\<close>
text (in limit_situation) \<open>
The assumption @{prop \<Phi>_in_terms_of_\<chi>} states that the component of @{term \<Phi>} at @{term a'}
is the arrow of @{term[source=true] S} corresponding to the function that takes an arrow
@{term "f \<in> C.hom a' a"} and produces the cone with vertex @{term a'} obtained
by transforming the universal cone @{term \<chi>} by @{term f}.
\<close>
subsection "Limit Cones Induce Limit Situations"
text\<open>
To obtain a limit situation from a limit cone, we need to introduce a set category
that is large enough to contain the hom-sets of @{term C} as well as the cones
over @{term D}. We use the category of all @{typ "('c + ('j \<Rightarrow> 'c))"}-sets for this.
\<close>
context limit_cone
begin
interpretation Cop: dual_category C ..
interpretation CopxC: product_category Cop.comp C ..
interpretation S: replete_setcat \<open>undefined :: 'c + ('j \<Rightarrow> 'c)\<close> .
notation S.comp (infixr "\<cdot>\<^sub>S" 55)
interpretation Sr: replete_concrete_set_category S.comp UNIV \<open>S.UP o Inr\<close>
apply unfold_locales
using S.UP_mapsto
apply auto[1]
using S.inj_UP inj_Inr inj_compose
by metis
interpretation Cones: cones_functor J C D S.comp \<open>S.UP o Inr\<close> ..
interpretation Hom: hom_functor C S.comp S.setp \<open>\<lambda>_. S.UP o Inl\<close>
apply (unfold_locales)
using S.UP_mapsto
apply auto[1]
using S.inj_UP injD inj_onI inj_Inl inj_compose
apply (metis (no_types, lifting))
using S.UP_mapsto
by auto
interpretation Y: yoneda_functor C S.comp S.setp \<open>\<lambda>_. S.UP o Inl\<close> ..
interpretation Ya: yoneda_functor_fixed_object C S.comp S.setp \<open>\<lambda>_. S.UP o Inl\<close> a
apply (unfold_locales) using ide_apex by auto
abbreviation inl :: "'c \<Rightarrow> 'c + ('j \<Rightarrow> 'c)" where "inl \<equiv> Inl"
abbreviation inr :: "('j \<Rightarrow> 'c) \<Rightarrow> 'c + ('j \<Rightarrow> 'c)" where "inr \<equiv> Inr"
abbreviation \<iota> where "\<iota> \<equiv> S.UP o inr"
abbreviation \<o> where "\<o> \<equiv> Cones.\<o>"
abbreviation \<phi> where "\<phi> \<equiv> \<lambda>_. S.UP o inl"
abbreviation \<psi> where "\<psi> \<equiv> Hom.\<psi>"
abbreviation Y where "Y \<equiv> Y.Y"
lemma Ya_ide:
assumes a': "C.ide a'"
shows "Y a a' = S.mkIde (Hom.set (a', a))"
using assms ide_apex Y.Y_simp Hom.map_ide by simp
lemma Ya_arr:
assumes g: "C.arr g"
shows "Y a g = S.mkArr (Hom.set (C.cod g, a)) (Hom.set (C.dom g, a))
(\<phi> (C.dom g, a) o Cop.comp g o \<psi> (C.cod g, a))"
using ide_apex g Y.Y_ide_arr [of a g "C.dom g" "C.cod g"] by auto
lemma cone_\<chi> [simp]:
shows "\<chi> \<in> D.cones a"
using cone_axioms by simp
text\<open>
For each object @{term a'} of @{term[source=true] C} we have a function mapping
@{term "C.hom a' a"} to the set of cones over @{term D} with apex @{term a'},
which takes @{term "f \<in> C.hom a' a"} to \<open>\<chi>f\<close>, where \<open>\<chi>f\<close> is the cone obtained by
composing @{term \<chi>} with @{term f} (after accounting for coercions to and from the
universe of @{term S}). The corresponding arrows of @{term S} are the
components of a natural isomorphism from @{term "Y a"} to \<open>Cones\<close>.
\<close>
definition \<Phi>o :: "'c \<Rightarrow> ('c + ('j \<Rightarrow> 'c)) setcat.arr"
where
"\<Phi>o a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>))"
lemma \<Phi>o_in_hom:
assumes a': "C.ide a'"
shows "\<guillemotleft>\<Phi>o a' : S.mkIde (Hom.set (a', a)) \<rightarrow>\<^sub>S S.mkIde (\<iota> ` D.cones a')\<guillemotright>"
proof -
have " \<guillemotleft>S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)) :
S.mkIde (Hom.set (a', a)) \<rightarrow>\<^sub>S S.mkIde (\<iota> ` D.cones a')\<guillemotright>"
proof -
have "(\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)) \<in> Hom.set (a', a) \<rightarrow> \<iota> ` D.cones a'"
proof
fix x
assume x: "x \<in> Hom.set (a', a)"
hence "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>"
using ide_apex a' Hom.\<psi>_mapsto by auto
hence "D.cones_map (\<psi> (a', a) x) \<chi> \<in> D.cones a'"
using ide_apex a' x D.cones_map_mapsto cone_\<chi> by force
thus "\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>) \<in> \<iota> ` D.cones a'" by simp
qed
moreover have "Hom.set (a', a) \<subseteq> S.Univ"
using ide_apex a' Hom.set_subset_Univ by auto
moreover have "\<iota> ` D.cones a' \<subseteq> S.Univ"
using S.UP_mapsto by auto
ultimately show ?thesis using S.mkArr_in_hom by simp
qed
thus ?thesis using \<Phi>o_def [of a'] by auto
qed
interpretation \<Phi>: transformation_by_components Cop.comp S.comp \<open>Y a\<close> Cones.map \<Phi>o
proof
fix a'
assume A': "Cop.ide a'"
show "\<guillemotleft>\<Phi>o a' : Y a a' \<rightarrow>\<^sub>S Cones.map a'\<guillemotright>"
using A' Ya_ide \<Phi>o_in_hom Cones.map_ide by auto
next
fix g
assume g: "Cop.arr g"
show "\<Phi>o (Cop.cod g) \<cdot>\<^sub>S Y a g = Cones.map g \<cdot>\<^sub>S \<Phi>o (Cop.dom g)"
proof -
let ?A = "Hom.set (C.cod g, a)"
let ?B = "Hom.set (C.dom g, a)"
let ?B' = "\<iota> ` D.cones (C.cod g)"
let ?C = "\<iota> ` D.cones (C.dom g)"
let ?F = "\<phi> (C.dom g, a) o Cop.comp g o \<psi> (C.cod g, a)"
let ?F' = "\<iota> o D.cones_map g o \<o>"
let ?G = "\<lambda>x. \<iota> (D.cones_map (\<psi> (C.dom g, a) x) \<chi>)"
let ?G' = "\<lambda>x. \<iota> (D.cones_map (\<psi> (C.cod g, a) x) \<chi>)"
have "S.arr (Y a g) \<and> Y a g = S.mkArr ?A ?B ?F"
using ide_apex g Ya.preserves_arr Ya_arr by fastforce
moreover have "S.arr (\<Phi>o (Cop.cod g))"
using g \<Phi>o_in_hom [of "Cop.cod g"] by auto
moreover have "\<Phi>o (Cop.cod g) = S.mkArr ?B ?C ?G"
using g \<Phi>o_def [of "C.dom g"] by auto
moreover have "S.seq (\<Phi>o (Cop.cod g)) (Y a g)"
using ide_apex g \<Phi>o_in_hom [of "Cop.cod g"] by auto
ultimately have 1: "S.seq (\<Phi>o (Cop.cod g)) (Y a g) \<and>
\<Phi>o (Cop.cod g) \<cdot>\<^sub>S Y a g = S.mkArr ?A ?C (?G o ?F)"
using S.comp_mkArr [of ?A ?B ?F ?C ?G] by argo
have "Cones.map g = S.mkArr (\<iota> ` D.cones (C.cod g)) (\<iota> ` D.cones (C.dom g)) ?F'"
using g Cones.map_simp by fastforce
moreover have "\<Phi>o (Cop.dom g) = S.mkArr ?A ?B' ?G'"
using g \<Phi>o_def by fastforce
moreover have "S.seq (Cones.map g) (\<Phi>o (Cop.dom g))"
using g Cones.preserves_hom [of g "C.cod g" "C.dom g"] \<Phi>o_in_hom [of "Cop.dom g"]
by force
ultimately have
2: "S.seq (Cones.map g) (\<Phi>o (Cop.dom g)) \<and>
Cones.map g \<cdot>\<^sub>S \<Phi>o (Cop.dom g) = S.mkArr ?A ?C (?F' o ?G')"
using S.seqI' [of "\<Phi>o (Cop.dom g)" "Cones.map g"] S.comp_mkArr by auto
have "\<Phi>o (Cop.cod g) \<cdot>\<^sub>S Y a g = S.mkArr ?A ?C (?G o ?F)"
using 1 by auto
also have "... = S.mkArr ?A ?C (?F' o ?G')"
proof (intro S.mkArr_eqI')
show "S.arr (S.mkArr ?A ?C (?G o ?F))" using 1 by force
show "\<And>x. x \<in> ?A \<Longrightarrow> (?G o ?F) x = (?F' o ?G') x"
proof -
fix x
assume x: "x \<in> ?A"
hence 1: "\<guillemotleft>\<psi> (C.cod g, a) x : C.cod g \<rightarrow> a\<guillemotright>"
using ide_apex g Hom.\<psi>_mapsto [of "C.cod g" a] by auto
have "(?G o ?F) x = \<iota> (D.cones_map (\<psi> (C.dom g, a)
(\<phi> (C.dom g, a) (\<psi> (C.cod g, a) x \<cdot> g))) \<chi>)"
proof - (* Why is it so balky with this proof? *)
have "(?G o ?F) x = ?G (?F x)" by simp
also have "... = \<iota> (D.cones_map (\<psi> (C.dom g, a)
(\<phi> (C.dom g, a) (\<psi> (C.cod g, a) x \<cdot> g))) \<chi>)"
by (metis Cop.comp_def comp_apply)
finally show ?thesis by auto
qed
also have "... = \<iota> (D.cones_map (\<psi> (C.cod g, a) x \<cdot> g) \<chi>)"
proof -
have "\<guillemotleft>\<psi> (C.cod g, a) x \<cdot> g : C.dom g \<rightarrow> a\<guillemotright>" using g 1 by auto
thus ?thesis using Hom.\<psi>_\<phi> by presburger
qed
also have "... = \<iota> (D.cones_map g (D.cones_map (\<psi> (C.cod g, a) x) \<chi>))"
using g x 1 cone_\<chi> D.cones_map_comp [of "\<psi> (C.cod g, a) x" g] by fastforce
also have "... = \<iota> (D.cones_map g (\<o> (\<iota> (D.cones_map (\<psi> (C.cod g, a) x) \<chi>))))"
using 1 cone_\<chi> D.cones_map_mapsto Sr.DN_UP by auto
also have "... = (?F' o ?G') x" by simp
finally show "(?G o ?F) x = (?F' o ?G') x" by auto
qed
qed
also have "... = Cones.map g \<cdot>\<^sub>S \<Phi>o (Cop.dom g)"
using 2 by auto
finally show ?thesis by auto
qed
qed
interpretation \<Phi>: set_valued_transformation Cop.comp S.comp S.setp
\<open>Y a\<close> Cones.map \<Phi>.map ..
interpretation \<Phi>: natural_isomorphism Cop.comp S.comp \<open>Y a\<close> Cones.map \<Phi>.map
proof
fix a'
assume a': "Cop.ide a'"
show "S.iso (\<Phi>.map a')"
proof -
let ?F = "\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
have bij: "bij_betw ?F (Hom.set (a', a)) (\<iota> ` D.cones a')"
proof -
have "\<And>x x'. \<lbrakk> x \<in> Hom.set (a', a); x' \<in> Hom.set (a', a);
\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>) \<rbrakk>
\<Longrightarrow> x = x'"
proof -
fix x x'
assume x: "x \<in> Hom.set (a', a)" and x': "x' \<in> Hom.set (a', a)"
and xx': "\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>)"
have \<psi>x: "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>" using x ide_apex a' Hom.\<psi>_mapsto by auto
have \<psi>x': "\<guillemotleft>\<psi> (a', a) x' : a' \<rightarrow> a\<guillemotright>" using x' ide_apex a' Hom.\<psi>_mapsto by auto
have 1: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof -
have "D.cones_map (\<psi> (a', a) x) \<chi> \<in> D.cones a'"
using \<psi>x a' cone_\<chi> D.cones_map_mapsto by force
hence 2: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>"
using a' is_universal by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof -
have "\<And>f. \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)
\<longleftrightarrow> D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>"
proof -
fix f :: 'c
have "D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>
\<longrightarrow> \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
by simp
thus "(\<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>))
= (D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>)"
by (meson Sr.inj_UP injD)
qed
thus ?thesis using 2 by auto
qed
qed
have 2: "\<exists>!x''. x'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof -
from 1 obtain f'' where
f'': "\<guillemotleft>f'' : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f'' \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
by blast
have "\<phi> (a', a) f'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) (\<phi> (a', a) f'')) \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof
show "\<phi> (a', a) f'' \<in> Hom.set (a', a)" using f'' Hom.set_def by auto
show "\<iota> (D.cones_map (\<psi> (a', a) (\<phi> (a', a) f'')) \<chi>) =
\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
using f'' Hom.\<psi>_\<phi> by presburger
qed
moreover have
"\<And>x''. x'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)
\<Longrightarrow> x'' = \<phi> (a', a) f''"
proof -
fix x''
assume x'': "x'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
hence "\<guillemotleft>\<psi> (a', a) x'' : a' \<rightarrow> a\<guillemotright> \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
using ide_apex a' Hom.set_def Hom.\<psi>_mapsto [of a' a] by auto
hence "\<phi> (a', a) (\<psi> (a', a) x'') = \<phi> (a', a) f''"
using 1 f'' by auto
thus "x'' = \<phi> (a', a) f''"
using ide_apex a' x'' Hom.\<phi>_\<psi> by simp
qed
ultimately show ?thesis
using ex1I [of "\<lambda>x'. x' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x') \<chi>) =
\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
"\<phi> (a', a) f''"]
by simp
qed
thus "x = x'" using x x' xx' by auto
qed
hence "inj_on ?F (Hom.set (a', a))"
using inj_onI [of "Hom.set (a', a)" ?F] by auto
moreover have "?F ` Hom.set (a', a) = \<iota> ` D.cones a'"
proof
show "?F ` Hom.set (a', a) \<subseteq> \<iota> ` D.cones a'"
proof
fix X'
assume X': "X' \<in> ?F ` Hom.set (a', a)"
from this obtain x' where x': "x' \<in> Hom.set (a', a) \<and> ?F x' = X'" by blast
show "X' \<in> \<iota> ` D.cones a'"
proof -
have "X' = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>)" using x' by blast
hence "X' = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>)" using x' by force
moreover have "\<guillemotleft>\<psi> (a', a) x' : a' \<rightarrow> a\<guillemotright>"
using ide_apex a' x' Hom.set_def Hom.\<psi>_\<phi> by auto
ultimately show ?thesis
using x' cone_\<chi> D.cones_map_mapsto by force
qed
qed
show "\<iota> ` D.cones a' \<subseteq> ?F ` Hom.set (a', a)"
proof
fix X'
assume X': "X' \<in> \<iota> ` D.cones a'"
hence "\<o> X' \<in> \<o> ` \<iota> ` D.cones a'" by simp
with Sr.DN_UP have "\<o> X' \<in> D.cones a'"
by auto
hence "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<o> X'"
using a' is_universal by simp
from this obtain f where "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<o> X'"
by auto
hence f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f \<chi>) = X'"
using X' Sr.UP_DN by auto
have "X' = ?F (\<phi> (a', a) f)"
using f Hom.\<psi>_\<phi> by presburger
thus "X' \<in> ?F ` Hom.set (a', a)"
using f Hom.set_def by force
qed
qed
ultimately show ?thesis
using bij_betw_def [of ?F "Hom.set (a', a)" "\<iota> ` D.cones a'"] inj_on_def by auto
qed
let ?f = "S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F"
have iso: "S.iso ?f"
proof -
have "?F \<in> Hom.set (a', a) \<rightarrow> \<iota> ` D.cones a'"
using bij bij_betw_imp_funcset by fast
hence 1: "S.arr ?f"
using ide_apex a' Hom.set_subset_Univ S.UP_mapsto by auto
thus ?thesis using bij S.iso_char S.set_mkIde by fastforce
qed
moreover have "?f = \<Phi>.map a'"
using a' \<Phi>o_def by force
finally show ?thesis by auto
qed
qed
interpretation R: representation_of_functor C S.comp S.setp \<phi> Cones.map a \<Phi>.map ..
lemma \<chi>_in_terms_of_\<Phi>:
shows "\<chi> = \<o> (\<Phi>.FUN a (\<phi> (a, a) a))"
proof -
have "\<Phi>.FUN a (\<phi> (a, a) a) =
(\<lambda>x \<in> Hom.set (a, a). \<iota> (D.cones_map (\<psi> (a, a) x) \<chi>)) (\<phi> (a, a) a)"
using ide_apex S.Fun_mkArr \<Phi>.map_simp_ide \<Phi>o_def \<Phi>.preserves_reflects_arr [of a]
by simp
also have "... = \<iota> (D.cones_map a \<chi>)"
proof -
have "(\<lambda>x \<in> Hom.set (a, a). \<iota> (D.cones_map (\<psi> (a, a) x) \<chi>)) (\<phi> (a, a) a)
= \<iota> (D.cones_map (\<psi> (a, a) (\<phi> (a, a) a)) \<chi>)"
proof -
have "\<phi> (a, a) a \<in> Hom.set (a, a)"
using ide_apex Hom.\<phi>_mapsto by fastforce
thus ?thesis
using restrict_apply' [of "\<phi> (a, a) a" "Hom.set (a, a)"] by blast
qed
also have "... = \<iota> (D.cones_map a \<chi>)"
proof -
have "\<psi> (a, a) (\<phi> (a, a) a) = a"
using ide_apex Hom.\<psi>_\<phi> [of a a a] by fastforce
thus ?thesis by metis
qed
finally show ?thesis by auto
qed
finally have "\<Phi>.FUN a (\<phi> (a, a) a) = \<iota> (D.cones_map a \<chi>)" by auto
also have "... = \<iota> \<chi>"
using ide_apex D.cones_map_ide [of \<chi> a] cone_\<chi> by simp
finally have "\<Phi>.FUN a (\<phi> (a, a) a) = \<iota> \<chi>" by blast
hence "\<o> (\<Phi>.FUN a (\<phi> (a, a) a)) = \<o> (\<iota> \<chi>)" by simp
thus ?thesis using cone_\<chi> Sr.DN_UP by simp
qed
abbreviation Hom
where "Hom \<equiv> Hom.map"
abbreviation \<Phi>
where "\<Phi> \<equiv> \<Phi>.map"
lemma induces_limit_situation:
shows "limit_situation J C D S.comp \<phi> \<iota> a \<Phi> \<chi>"
using \<chi>_in_terms_of_\<Phi> \<Phi>o_def by unfold_locales auto
no_notation S.comp (infixr "\<cdot>\<^sub>S" 55)
end
sublocale limit_cone \<subseteq> limit_situation J C D replete_setcat.comp \<phi> \<iota> a \<Phi> \<chi>
using induces_limit_situation by auto
subsection "Representations of the Cones Functor Induce Limit Situations"
context representation_of_cones_functor
begin
interpretation \<Phi>: set_valued_transformation Cop.comp S S.setp \<open>Y a\<close> Cones.map \<Phi> ..
interpretation \<Psi>: inverse_transformation Cop.comp S \<open>Y a\<close> Cones.map \<Phi> ..
interpretation \<Psi>: set_valued_transformation Cop.comp S S.setp
Cones.map \<open>Y a\<close> \<Psi>.map ..
abbreviation \<o>
where "\<o> \<equiv> Cones.\<o>"
abbreviation \<chi>
where "\<chi> \<equiv> \<o> (S.Fun (\<Phi> a) (\<phi> (a, a) a))"
lemma Cones_SET_eq_\<iota>_img_cones:
assumes "C.ide a'"
shows "Cones.SET a' = \<iota> ` D.cones a'"
proof -
have "\<iota> ` D.cones a' \<subseteq> S.Univ" using S.UP_mapsto by auto
thus ?thesis using assms Cones.map_ide S.set_mkIde by auto
qed
lemma \<iota>\<chi>:
shows "\<iota> \<chi> = S.Fun (\<Phi> a) (\<phi> (a, a) a)"
proof -
have "S.Fun (\<Phi> a) (\<phi> (a, a) a) \<in> Cones.SET a"
using Ya.ide_a Hom.\<phi>_mapsto S.Fun_mapsto [of "\<Phi> a"] Hom.set_map by fastforce
thus ?thesis
using Ya.ide_a Cones_SET_eq_\<iota>_img_cones by auto
qed
interpretation \<chi>: cone J C D a \<chi>
proof -
have "\<iota> \<chi> \<in> \<iota> ` D.cones a"
using Ya.ide_a \<iota>\<chi> S.Fun_mapsto [of "\<Phi> a"] Hom.\<phi>_mapsto Hom.set_map
Cones_SET_eq_\<iota>_img_cones by fastforce
thus "D.cone a \<chi>"
by (metis (no_types, lifting) S.DN_UP UNIV_I f_inv_into_f inv_into_into mem_Collect_eq)
qed
lemma cone_\<chi>:
shows "D.cone a \<chi>" ..
lemma \<Phi>_FUN_simp:
assumes a': "C.ide a'" and x: "x \<in> Hom.set (a', a)"
shows "\<Phi>.FUN a' x = Cones.FUN (\<psi> (a', a) x) (\<iota> \<chi>)"
proof -
have \<psi>x: "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>"
using Ya.ide_a a' x Hom.\<psi>_mapsto by blast
have \<phi>a: "\<phi> (a, a) a \<in> Hom.set (a, a)" using Ya.ide_a Hom.\<phi>_mapsto by fastforce
have "\<Phi>.FUN a' x = (\<Phi>.FUN a' o Ya.FUN (\<psi> (a', a) x)) (\<phi> (a, a) a)"
proof -
have "\<phi> (a', a) (a \<cdot> \<psi> (a', a) x) = x"
using Ya.ide_a a' x \<psi>x Hom.\<phi>_\<psi> C.comp_cod_arr by fastforce
moreover have "S.arr (S.mkArr (Hom.set (a, a)) (Hom.set (a', a))
(\<phi> (a', a) \<circ> Cop.comp (\<psi> (a', a) x) \<circ> \<psi> (a, a)))"
by (metis C.arrI Cop.arr_char Ya.Y_ide_arr(2) Ya.preserves_arr \<chi>.ide_apex \<psi>x)
ultimately show ?thesis
using Ya.ide_a a' x Ya.Y_ide_arr \<psi>x \<phi>a C.ide_in_hom by auto
qed
also have "... = (Cones.FUN (\<psi> (a', a) x) o \<Phi>.FUN a) (\<phi> (a, a) a)"
proof -
have "(\<Phi>.FUN a' o Ya.FUN (\<psi> (a', a) x)) (\<phi> (a, a) a)
= S.Fun (\<Phi> a' \<cdot>\<^sub>S Y a (\<psi> (a', a) x)) (\<phi> (a, a) a)"
using \<psi>x a' \<phi>a Ya.ide_a Ya.map_simp Hom.set_map by (elim C.in_homE, auto)
also have "... = S.Fun (S (Cones.map (\<psi> (a', a) x)) (\<Phi> a)) (\<phi> (a, a) a)"
using \<psi>x is_natural_1 [of "\<psi> (a', a) x"] is_natural_2 [of "\<psi> (a', a) x"] by auto
also have "... = (Cones.FUN (\<psi> (a', a) x) o \<Phi>.FUN a) (\<phi> (a, a) a)"
proof -
have "S.seq (Cones.map (\<psi> (a', a) x)) (\<Phi> a)"
using Ya.ide_a \<psi>x Cones.map_preserves_dom [of "\<psi> (a', a) x"]
apply (intro S.seqI)
apply auto[2]
by fastforce
thus ?thesis
using Ya.ide_a \<phi>a Hom.set_map by auto
qed
finally show ?thesis by simp
qed
also have "... = Cones.FUN (\<psi> (a', a) x) (\<iota> \<chi>)" using \<iota>\<chi> by simp
finally show ?thesis by auto
qed
lemma \<chi>_is_universal:
assumes "D.cone a' \<chi>'"
shows "\<guillemotleft>\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>')) : a' \<rightarrow> a\<guillemotright>"
and "D.cones_map (\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))) \<chi> = \<chi>'"
and "\<lbrakk> \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>; D.cones_map f' \<chi> = \<chi>' \<rbrakk> \<Longrightarrow> f' = \<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
proof -
interpret \<chi>': cone J C D a' \<chi>' using assms by auto
have a': "C.ide a'" using \<chi>'.ide_apex by simp
have \<iota>\<chi>': "\<iota> \<chi>' \<in> Cones.SET a'" using assms a' Cones_SET_eq_\<iota>_img_cones by auto
let ?f = "\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
have A: "\<Psi>.FUN a' (\<iota> \<chi>') \<in> Hom.set (a', a)"
proof -
have "\<Psi>.FUN a' \<in> Cones.SET a' \<rightarrow> Ya.SET a'"
using a' \<Psi>.preserves_hom [of a' a' a'] S.Fun_mapsto [of "\<Psi>.map a'"] by fastforce
thus ?thesis using a' \<iota>\<chi>' Ya.ide_a Hom.set_map by auto
qed
show f: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright>" using A a' Ya.ide_a Hom.\<psi>_mapsto [of a' a] by auto
have E: "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> Cones.FUN f (\<iota> \<chi>) = \<Phi>.FUN a' (\<phi> (a', a) f)"
proof -
fix f
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>"
have "\<phi> (a', a) f \<in> Hom.set (a', a)"
using a' Ya.ide_a f Hom.\<phi>_mapsto by auto
thus "Cones.FUN f (\<iota> \<chi>) = \<Phi>.FUN a' (\<phi> (a', a) f)"
using a' f \<Phi>_FUN_simp by simp
qed
have I: "\<Phi>.FUN a' (\<Psi>.FUN a' (\<iota> \<chi>')) = \<iota> \<chi>'"
proof -
have "\<Phi>.FUN a' (\<Psi>.FUN a' (\<iota> \<chi>')) =
compose (\<Psi>.DOM a') (\<Phi>.FUN a') (\<Psi>.FUN a') (\<iota> \<chi>')"
using a' \<iota>\<chi>' Cones.map_ide \<Psi>.preserves_hom [of a' a' a'] by force
also have "... = (\<lambda>x \<in> \<Psi>.DOM a'. x) (\<iota> \<chi>')"
using a' \<Psi>.inverts_components S.inverse_arrows_char by force
also have "... = \<iota> \<chi>'"
using a' \<iota>\<chi>' Cones.map_ide \<Psi>.preserves_hom [of a' a' a'] by force
finally show ?thesis by auto
qed
show f\<chi>: "D.cones_map ?f \<chi> = \<chi>'"
proof -
have "D.cones_map ?f \<chi> = (\<o> o Cones.FUN ?f o \<iota>) \<chi>"
using f Cones.preserves_arr [of ?f] cone_\<chi>
by (cases "D.cone a \<chi>", auto)
also have "... = \<chi>'"
using f Ya.ide_a a' A E I by auto
finally show ?thesis by auto
qed
show "\<lbrakk> \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>; D.cones_map f' \<chi> = \<chi>' \<rbrakk> \<Longrightarrow> f' = ?f"
proof -
assume f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>" and f'\<chi>: "D.cones_map f' \<chi> = \<chi>'"
show "f' = ?f"
proof -
have 1: "\<phi> (a', a) f' \<in> Hom.set (a', a) \<and> \<phi> (a', a) ?f \<in> Hom.set (a', a)"
using Ya.ide_a a' f f' Hom.\<phi>_mapsto by auto
have "S.iso (\<Phi> a')" using \<chi>'.ide_apex components_are_iso by auto
hence 2: "S.arr (\<Phi> a') \<and> bij_betw (\<Phi>.FUN a') (Hom.set (a', a)) (Cones.SET a')"
using Ya.ide_a a' S.iso_char Hom.set_map by auto
have "\<Phi>.FUN a' (\<phi> (a', a) f') = \<Phi>.FUN a' (\<phi> (a', a) ?f)"
proof -
have "\<Phi>.FUN a' (\<phi> (a', a) ?f) = \<iota> \<chi>'"
using A I Hom.\<phi>_\<psi> Ya.ide_a a' by simp
also have "... = Cones.FUN f' (\<iota> \<chi>)"
using f f' A E cone_\<chi> Cones.preserves_arr f\<chi> f'\<chi> by (elim C.in_homE, auto)
also have "... = \<Phi>.FUN a' (\<phi> (a', a) f')"
using f' E by simp
finally show ?thesis by argo
qed
moreover have "inj_on (\<Phi>.FUN a') (Hom.set (a', a))"
using 2 bij_betw_imp_inj_on by blast
ultimately have 3: "\<phi> (a', a) f' = \<phi> (a', a) ?f"
using 1 inj_on_def [of "\<Phi>.FUN a'" "Hom.set (a', a)"] by blast
show ?thesis
proof -
have "f' = \<psi> (a', a) (\<phi> (a', a) f')"
using Ya.ide_a a' f' Hom.\<psi>_\<phi> by simp
also have "... = \<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
using Ya.ide_a a' Hom.\<psi>_\<phi> A 3 by simp
finally show ?thesis by blast
qed
qed
qed
qed
interpretation \<chi>: limit_cone J C D a \<chi>
proof
show "\<And>a' \<chi>'. D.cone a' \<chi>' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
proof -
fix a' \<chi>'
assume 1: "D.cone a' \<chi>'"
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
proof
show "\<guillemotleft>\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>')) : a' \<rightarrow> a\<guillemotright> \<and>
D.cones_map (\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))) \<chi> = \<chi>'"
using 1 \<chi>_is_universal by blast
show "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' \<Longrightarrow> f = \<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
using 1 \<chi>_is_universal by blast
qed
qed
qed
lemma \<chi>_is_limit_cone:
shows "D.limit_cone a \<chi>" ..
lemma induces_limit_situation:
shows "limit_situation J C D S \<phi> \<iota> a \<Phi> \<chi>"
proof
show "\<chi> = \<chi>" by simp
fix a'
assume a': "Cop.ide a'"
let ?F = "\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
show "\<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F"
proof -
have 1: "\<guillemotleft>\<Phi> a' : S.mkIde (Hom.set (a', a)) \<rightarrow>\<^sub>S S.mkIde (\<iota> ` D.cones a')\<guillemotright>"
using a' Cones.map_ide Ya.ide_a by auto
moreover have "\<Phi>.DOM a' = Hom.set (a', a)"
using 1 Hom.set_subset_Univ a' Ya.ide_a Hom.set_map by simp
moreover have "\<Phi>.COD a' = \<iota> ` D.cones a'"
using a' Cones_SET_eq_\<iota>_img_cones by fastforce
ultimately have 2: "\<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<Phi>.FUN a')"
using S.mkArr_Fun [of "\<Phi> a'"] by fastforce
also have "... = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F"
proof
show "S.arr (S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<Phi>.FUN a'))"
using 1 2 by auto
show "\<And>x. x \<in> Hom.set (a', a) \<Longrightarrow> \<Phi>.FUN a' x = ?F x"
proof -
fix x
assume x: "x \<in> Hom.set (a', a)"
hence \<psi>x: "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>"
using a' Ya.ide_a Hom.\<psi>_mapsto by auto
show "\<Phi>.FUN a' x = ?F x"
proof -
have "\<Phi>.FUN a' x = Cones.FUN (\<psi> (a', a) x) (\<iota> \<chi>)"
using a' x \<Phi>_FUN_simp by simp
also have "... = restrict (\<iota> o D.cones_map (\<psi> (a', a) x) o \<o>) (\<iota> ` D.cones a) (\<iota> \<chi>)"
using \<psi>x Cones.map_simp Cones.preserves_arr [of "\<psi> (a', a) x"] S.Fun_mkArr
by (elim C.in_homE, auto)
also have "... = ?F x" using cone_\<chi> by simp
ultimately show ?thesis by simp
qed
qed
qed
finally show "\<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F" by auto
qed
qed
end
sublocale representation_of_cones_functor \<subseteq> limit_situation J C D S \<phi> \<iota> a \<Phi> \<chi>
using induces_limit_situation by auto
section "Categories with Limits"
context category
begin
text\<open>
A category @{term[source=true] C} has limits of shape @{term J} if every diagram of shape
@{term J} admits a limit cone.
\<close>
definition has_limits_of_shape
where "has_limits_of_shape J \<equiv> \<forall>D. diagram J C D \<longrightarrow> (\<exists>a \<chi>. limit_cone J C D a \<chi>)"
text\<open>
A category has limits at a type @{typ 'j} if it has limits of shape @{term J}
for every category @{term J} whose arrows are of type @{typ 'j}.
\<close>
definition has_limits
where "has_limits (_ :: 'j) \<equiv> \<forall>J :: 'j comp. category J \<longrightarrow> has_limits_of_shape J"
text\<open>
Whether a category has limits of shape \<open>J\<close> truly depends only on the ``shape''
(\emph{i.e.}~isomorphism class) of \<open>J\<close> and not on details of its construction.
\<close>
lemma has_limits_preserved_by_isomorphism:
assumes "has_limits_of_shape J" and "isomorphic_categories J J'"
shows "has_limits_of_shape J'"
proof -
interpret J: category J
using assms(2) isomorphic_categories_def isomorphic_categories_axioms_def by auto
interpret J': category J'
using assms(2) isomorphic_categories_def isomorphic_categories_axioms_def by auto
from assms(2) obtain \<phi> \<psi> where IF: "inverse_functors J' J \<phi> \<psi>"
using isomorphic_categories_def isomorphic_categories_axioms_def
inverse_functors_sym
by blast
interpret IF: inverse_functors J' J \<phi> \<psi> using IF by auto
have \<psi>\<phi>: "\<psi> o \<phi> = J.map" using IF.inv by metis
have \<phi>\<psi>: "\<phi> o \<psi> = J'.map" using IF.inv' by metis
have "\<And>D'. diagram J' C D' \<Longrightarrow> \<exists>a \<chi>. limit_cone J' C D' a \<chi>"
proof -
fix D'
assume D': "diagram J' C D'"
interpret D': diagram J' C D' using D' by auto
interpret D: composite_functor J J' C \<phi> D' ..
interpret D: diagram J C \<open>D' o \<phi>\<close> ..
have D: "diagram J C (D' o \<phi>)" ..
from assms(1) obtain a \<chi> where \<chi>: "D.limit_cone a \<chi>"
using D has_limits_of_shape_def by blast
interpret \<chi>: limit_cone J C \<open>D' o \<phi>\<close> a \<chi> using \<chi> by auto
interpret A': constant_functor J' C a
using \<chi>.ide_apex by (unfold_locales, auto)
have \<chi>o\<psi>: "cone J' C (D' o \<phi> o \<psi>) a (\<chi> o \<psi>)"
using comp_cone_functor IF.G.functor_axioms \<chi>.cone_axioms by fastforce
hence \<chi>o\<psi>: "cone J' C D' a (\<chi> o \<psi>)"
using \<phi>\<psi> by (metis D'.functor_axioms Fun.comp_assoc comp_functor_identity)
interpret \<chi>o\<psi>: cone J' C D' a \<open>\<chi> o \<psi>\<close> using \<chi>o\<psi> by auto
interpret \<chi>o\<psi>: limit_cone J' C D' a \<open>\<chi> o \<psi>\<close>
proof
fix a' \<chi>'
assume \<chi>': "D'.cone a' \<chi>'"
interpret \<chi>': cone J' C D' a' \<chi>' using \<chi>' by auto
have \<chi>'o\<phi>: "cone J C (D' o \<phi>) a' (\<chi>' o \<phi>)"
using \<chi>' comp_cone_functor IF.F.functor_axioms by fastforce
interpret \<chi>'o\<phi>: cone J C \<open>D' o \<phi>\<close> a' \<open>\<chi>' o \<phi>\<close> using \<chi>'o\<phi> by auto
have "cone J C (D' o \<phi>) a' (\<chi>' o \<phi>)" ..
hence 1: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' o \<phi>"
using \<chi>.is_universal by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f (\<chi> o \<psi>) = \<chi>'"
proof
let ?f = "THE f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' o \<phi>"
have f: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map ?f \<chi> = \<chi>' o \<phi>"
using 1 theI' [of "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' o \<phi>"] by blast
have f_in_hom: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright>" using f by blast
have "D'.cones_map ?f (\<chi> o \<psi>) = \<chi>'"
proof
fix j'
have "\<not>J'.arr j' \<Longrightarrow> D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'"
proof -
assume j': "\<not>J'.arr j'"
have "D'.cones_map ?f (\<chi> o \<psi>) j' = null"
using j' f_in_hom \<chi>o\<psi> by fastforce
thus ?thesis
using j' \<chi>'.is_extensional by simp
qed
moreover have "J'.arr j' \<Longrightarrow> D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'"
proof -
assume j': "J'.arr j'"
have "D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi> (\<psi> j') \<cdot> ?f"
using j' f \<chi>o\<psi> by fastforce
also have "... = D.cones_map ?f \<chi> (\<psi> j')"
using j' f_in_hom \<chi> \<chi>.cone_\<chi> by fastforce
also have "... = \<chi>' j'"
using j' f \<chi> \<phi>\<psi> Fun.comp_def J'.map_simp by metis
finally show "D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'" by auto
qed
ultimately show "D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'" by blast
qed
thus "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map ?f (\<chi> o \<psi>) = \<chi>'" using f by auto
fix f'
assume f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f' (\<chi> o \<psi>) = \<chi>'"
have "D.cones_map f' \<chi> = \<chi>' o \<phi>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j"
using f' \<chi> \<chi>'o\<phi>.is_extensional \<chi>.cone_\<chi> mem_Collect_eq restrict_apply by auto
moreover have "J.arr j \<Longrightarrow> D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j"
proof -
assume j: "J.arr j"
have "D.cones_map f' \<chi> j = C (\<chi> j) f'"
using j f' \<chi>.cone_\<chi> by auto
also have "... = C ((\<chi> o \<psi>) (\<phi> j)) f'"
using j f' \<psi>\<phi> by (metis comp_apply J.map_simp)
also have "... = D'.cones_map f' (\<chi> o \<psi>) (\<phi> j)"
using j f' \<chi>o\<psi> by fastforce
also have "... = (\<chi>' o \<phi>) j"
using j f' by auto
finally show "D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j" by auto
qed
ultimately show "D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j" by blast
qed
hence "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f' \<chi> = \<chi>' o \<phi>"
using f' by auto
moreover have "\<And>P x x'. (\<exists>!x. P x) \<and> P x \<and> P x' \<Longrightarrow> x = x'"
by auto
ultimately show "f' = ?f" using 1 f by blast
qed
qed
have "limit_cone J' C D' a (\<chi> o \<psi>)" ..
thus "\<exists>a \<chi>. limit_cone J' C D' a \<chi>" by blast
qed
thus ?thesis using has_limits_of_shape_def by auto
qed
end
subsection "Diagonal Functors"
text\<open>
The existence of limits can also be expressed in terms of adjunctions: a category @{term C}
has limits of shape @{term J} if the diagonal functor taking each object @{term a}
in @{term C} to the constant-@{term a} diagram and each arrow \<open>f \<in> C.hom a a'\<close>
to the constant-@{term f} natural transformation between diagrams is a left adjoint functor.
\<close>
locale diagonal_functor =
C: category C +
J: category J +
J_C: functor_category J C
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
begin
notation J.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J _\<guillemotright>")
notation J_C.comp (infixr "\<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>]" 55)
notation J_C.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] _\<guillemotright>")
definition map :: "'c \<Rightarrow> ('j, 'c) J_C.arr"
where "map f = (if C.arr f then J_C.MkArr (constant_functor.map J C (C.dom f))
(constant_functor.map J C (C.cod f))
(constant_transformation.map J C f)
else J_C.null)"
lemma is_functor:
shows "functor C J_C.comp map"
proof
fix f
show "\<not> C.arr f \<Longrightarrow> local.map f = J_C.null"
using map_def by simp
assume f: "C.arr f"
interpret Dom_f: constant_functor J C \<open>C.dom f\<close>
using f by (unfold_locales, auto)
interpret Cod_f: constant_functor J C \<open>C.cod f\<close>
using f by (unfold_locales, auto)
interpret Fun_f: constant_transformation J C f
using f by (unfold_locales, auto)
show 1: "J_C.arr (map f)"
using f map_def by (simp add: Fun_f.natural_transformation_axioms)
show "J_C.dom (map f) = map (C.dom f)"
proof -
have "constant_transformation J C (C.dom f)"
using f by unfold_locales auto
hence "constant_transformation.map J C (C.dom f) = Dom_f.map"
using Dom_f.map_def constant_transformation.map_def [of J C "C.dom f"] by auto
thus ?thesis using f 1 by (simp add: map_def J_C.dom_char)
qed
show "J_C.cod (map f) = map (C.cod f)"
proof -
have "constant_transformation J C (C.cod f)"
using f by unfold_locales auto
hence "constant_transformation.map J C (C.cod f) = Cod_f.map"
using Cod_f.map_def constant_transformation.map_def [of J C "C.cod f"] by auto
thus ?thesis using f 1 by (simp add: map_def J_C.cod_char)
qed
next
fix f g
assume g: "C.seq g f"
have f: "C.arr f" using g by auto
interpret Dom_f: constant_functor J C \<open>C.dom f\<close>
using f by unfold_locales auto
interpret Cod_f: constant_functor J C \<open>C.cod f\<close>
using f by unfold_locales auto
interpret Fun_f: constant_transformation J C f
using f by unfold_locales auto
interpret Cod_g: constant_functor J C \<open>C.cod g\<close>
using g by unfold_locales auto
interpret Fun_g: constant_transformation J C g
using g by unfold_locales auto
interpret Fun_g: natural_transformation J C Cod_f.map Cod_g.map Fun_g.map
apply unfold_locales
using f g C.seqE [of g f] C.comp_arr_dom C.comp_cod_arr Fun_g.is_extensional by auto
interpret Fun_fg: vertical_composite
J C Dom_f.map Cod_f.map Cod_g.map Fun_f.map Fun_g.map ..
have 1: "J_C.arr (map f)"
using f map_def by (simp add: Fun_f.natural_transformation_axioms)
show "map (g \<cdot> f) = map g \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map f"
proof -
have "map (C g f) = J_C.MkArr Dom_f.map Cod_g.map
(constant_transformation.map J C (C g f))"
using f g map_def by simp
also have "... = J_C.MkArr Dom_f.map Cod_g.map (\<lambda>j. if J.arr j then C g f else C.null)"
proof -
have "constant_transformation J C (g \<cdot> f)"
using g by unfold_locales auto
thus ?thesis using constant_transformation.map_def by metis
qed
also have "... = J_C.comp (J_C.MkArr Cod_f.map Cod_g.map Fun_g.map)
(J_C.MkArr Dom_f.map Cod_f.map Fun_f.map)"
proof -
have "J_C.MkArr Cod_f.map Cod_g.map Fun_g.map \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>]
J_C.MkArr Dom_f.map Cod_f.map Fun_f.map
= J_C.MkArr Dom_f.map Cod_g.map Fun_fg.map"
using J_C.comp_char J_C.comp_MkArr Fun_f.natural_transformation_axioms
Fun_g.natural_transformation_axioms
by blast
also have "... = J_C.MkArr Dom_f.map Cod_g.map
(\<lambda>j. if J.arr j then g \<cdot> f else C.null)"
using Fun_fg.is_extensional Fun_fg.map_simp_2 by auto
finally show ?thesis by auto
qed
also have "... = map g \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map f"
using f g map_def by fastforce
finally show ?thesis by auto
qed
qed
sublocale "functor" C J_C.comp map
using is_functor by auto
text\<open>
The objects of \<open>[J, C]\<close> correspond bijectively to diagrams of shape @{term J}
in @{term C}.
\<close>
lemma ide_determines_diagram:
assumes "J_C.ide d"
shows "diagram J C (J_C.Map d)" and "J_C.MkIde (J_C.Map d) = d"
proof -
interpret \<delta>: natural_transformation J C \<open>J_C.Map d\<close> \<open>J_C.Map d\<close> \<open>J_C.Map d\<close>
using assms J_C.ide_char J_C.arr_MkArr by fastforce
interpret D: "functor" J C \<open>J_C.Map d\<close> ..
show "diagram J C (J_C.Map d)" ..
show "J_C.MkIde (J_C.Map d) = d"
using assms J_C.ide_char by (metis J_C.ideD(1) J_C.MkArr_Map)
qed
lemma diagram_determines_ide:
assumes "diagram J C D"
shows "J_C.ide (J_C.MkIde D)" and "J_C.Map (J_C.MkIde D) = D"
proof -
interpret D: diagram J C D using assms by auto
show "J_C.ide (J_C.MkIde D)" using J_C.ide_char
using D.functor_axioms J_C.ide_MkIde by auto
thus "J_C.Map (J_C.MkIde D) = D"
using J_C.in_homE by simp
qed
lemma bij_betw_ide_diagram:
shows "bij_betw J_C.Map (Collect J_C.ide) (Collect (diagram J C))"
proof (intro bij_betwI)
show "J_C.Map \<in> Collect J_C.ide \<rightarrow> Collect (diagram J C)"
using ide_determines_diagram by blast
show "J_C.MkIde \<in> Collect (diagram J C) \<rightarrow> Collect J_C.ide"
using diagram_determines_ide by blast
show "\<And>d. d \<in> Collect J_C.ide \<Longrightarrow> J_C.MkIde (J_C.Map d) = d"
using ide_determines_diagram by blast
show "\<And>D. D \<in> Collect (diagram J C) \<Longrightarrow> J_C.Map (J_C.MkIde D) = D"
using diagram_determines_ide by blast
qed
text\<open>
Arrows from from the diagonal functor correspond bijectively to cones.
\<close>
lemma arrow_determines_cone:
assumes "J_C.ide d" and "arrow_from_functor C J_C.comp map a d x"
shows "cone J C (J_C.Map d) a (J_C.Map x)"
and "J_C.MkArr (constant_functor.map J C a) (J_C.Map d) (J_C.Map x) = x"
proof -
interpret D: diagram J C \<open>J_C.Map d\<close>
using assms ide_determines_diagram by auto
interpret x: arrow_from_functor C J_C.comp map a d x
using assms by auto
interpret A: constant_functor J C a
using x.arrow by (unfold_locales, auto)
interpret \<alpha>: constant_transformation J C a
using x.arrow by (unfold_locales, auto)
have Dom_x: "J_C.Dom x = A.map"
using J_C.in_hom_char map_def x.arrow by force
have Cod_x: "J_C.Cod x = J_C.Map d"
using x.arrow by auto
interpret \<chi>: natural_transformation J C A.map \<open>J_C.Map d\<close> \<open>J_C.Map x\<close>
using x.arrow J_C.arr_char [of x] Dom_x Cod_x by force
show "D.cone a (J_C.Map x)" ..
show "J_C.MkArr A.map (J_C.Map d) (J_C.Map x) = x"
using x.arrow Dom_x Cod_x \<chi>.natural_transformation_axioms
by (intro J_C.arr_eqI, auto)
qed
lemma cone_determines_arrow:
assumes "J_C.ide d" and "cone J C (J_C.Map d) a \<chi>"
shows "arrow_from_functor C J_C.comp map a d
(J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \<chi>)"
and "J_C.Map (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \<chi>) = \<chi>"
proof -
interpret \<chi>: cone J C \<open>J_C.Map d\<close> a \<chi> using assms(2) by auto
let ?x = "J_C.MkArr \<chi>.A.map (J_C.Map d) \<chi>"
interpret x: arrow_from_functor C J_C.comp map a d ?x
proof
have "\<guillemotleft>J_C.MkArr \<chi>.A.map (J_C.Map d) \<chi> :
J_C.MkIde \<chi>.A.map \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde (J_C.Map d)\<guillemotright>"
using \<chi>.natural_transformation_axioms by auto
moreover have "J_C.MkIde \<chi>.A.map = map a"
using \<chi>.A.value_is_ide map_def \<chi>.A.map_def C.ide_char
by (metis (no_types, lifting) J_C.dom_MkArr preserves_arr preserves_dom)
moreover have "J_C.MkIde (J_C.Map d) = d"
using assms ide_determines_diagram(2) by simp
ultimately show "C.ide a \<and> \<guillemotleft>J_C.MkArr \<chi>.A.map (J_C.Map d) \<chi> : map a \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] d\<guillemotright>"
using \<chi>.A.value_is_ide by simp
qed
show "arrow_from_functor C J_C.comp map a d ?x" ..
show "J_C.Map (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \<chi>) = \<chi>"
by (simp add: \<chi>.natural_transformation_axioms)
qed
text\<open>
Transforming a cone by composing at the apex with an arrow @{term g} corresponds,
via the preceding bijections, to composition in \<open>[J, C]\<close> with the image of @{term g}
under the diagonal functor.
\<close>
lemma cones_map_is_composition:
assumes "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright>" and "cone J C D a \<chi>"
shows "J_C.MkArr (constant_functor.map J C a') D (diagram.cones_map J C D g \<chi>)
= J_C.MkArr (constant_functor.map J C a) D \<chi> \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g"
proof -
interpret A: constant_transformation J C a
using assms(1) by (unfold_locales, auto)
interpret \<chi>: cone J C D a \<chi> using assms(2) by auto
have cone_\<chi>: "cone J C D a \<chi>" ..
interpret A': constant_transformation J C a'
using assms(1) by (unfold_locales, auto)
let ?\<chi>' = "\<chi>.D.cones_map g \<chi>"
interpret \<chi>': cone J C D a' ?\<chi>'
using assms(1) cone_\<chi> \<chi>.D.cones_map_mapsto by blast
let ?x = "J_C.MkArr \<chi>.A.map D \<chi>"
let ?x' = "J_C.MkArr \<chi>'.A.map D ?\<chi>'"
show "?x' = J_C.comp ?x (map g)"
proof (intro J_C.arr_eqI)
have x: "J_C.arr ?x"
using \<chi>.natural_transformation_axioms J_C.arr_char [of ?x] by simp
show x': "J_C.arr ?x'"
using \<chi>'.natural_transformation_axioms J_C.arr_char [of ?x'] by simp
have 3: "\<guillemotleft>?x : map a \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde D\<guillemotright>"
using \<chi>.D.diagram_axioms arrow_from_functor.arrow cone_\<chi> cone_determines_arrow(1)
diagram_determines_ide(1)
by fastforce
have 4: "\<guillemotleft>?x' : map a' \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde D\<guillemotright>"
by (metis (no_types, lifting) J_C.Dom.simps(1) J_C.Dom_cod J_C.Map_cod
J_C.cod_MkArr \<chi>'.cone_axioms arrow_from_functor.arrow category.ide_cod
cone_determines_arrow(1) functor_def is_functor x)
have seq_xg: "J_C.seq ?x (map g)"
using assms(1) 3 preserves_hom [of g] by (intro J_C.seqI', auto)
show 2: "J_C.seq ?x (map g)"
using seq_xg J_C.seqI' by blast
show "J_C.Dom ?x' = J_C.Dom (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
proof -
have "J_C.Dom ?x' = J_C.Dom (J_C.dom ?x')"
using x' J_C.Dom_dom by simp
also have "... = J_C.Dom (map a')"
using 4 by force
also have "... = J_C.Dom (J_C.dom (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g))"
using assms(1) 2 by auto
also have "... = J_C.Dom (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
using seq_xg J_C.Dom_dom J_C.seqI' by blast
finally show ?thesis by auto
qed
show "J_C.Cod ?x' = J_C.Cod (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
proof -
have "J_C.Cod ?x' = J_C.Cod (J_C.cod ?x')"
using x' J_C.Cod_cod by simp
also have "... = J_C.Cod (J_C.MkIde D)"
using 4 by force
also have "... = J_C.Cod (J_C.cod (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g))"
using 2 3 J_C.cod_comp J_C.in_homE by metis
also have "... = J_C.Cod (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
using seq_xg J_C.Cod_cod J_C.seqI' by blast
finally show ?thesis by auto
qed
show "J_C.Map ?x' = J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
proof -
interpret g: constant_transformation J C g
apply unfold_locales using assms(1) by auto
interpret \<chi>og: vertical_composite J C A'.map \<chi>.A.map D g.map \<chi>
using assms(1) C.comp_arr_dom C.comp_cod_arr A'.is_extensional g.is_extensional
apply (unfold_locales, auto)
by (elim J.seqE, auto)
have "J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g) = \<chi>og.map"
using assms(1) 2 J_C.comp_char map_def by auto
also have "... = J_C.Map ?x'"
using x' \<chi>og.map_def J_C.arr_char [of ?x'] natural_transformation.is_extensional
assms(1) cone_\<chi> \<chi>og.map_simp_2
by fastforce
finally show ?thesis by auto
qed
qed
qed
text\<open>
Coextension along an arrow from a functor is equivalent to a transformation of cones.
\<close>
lemma coextension_iff_cones_map:
assumes x: "arrow_from_functor C J_C.comp map a d x"
and g: "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright>"
and x': "\<guillemotleft>x' : map a' \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] d\<guillemotright>"
shows "arrow_from_functor.is_coext C J_C.comp map a x a' x' g
\<longleftrightarrow> J_C.Map x' = diagram.cones_map J C (J_C.Map d) g (J_C.Map x)"
proof -
interpret x: arrow_from_functor C J_C.comp map a d x
using assms by auto
interpret A': constant_functor J C a'
using assms(2) by (unfold_locales, auto)
have x': "arrow_from_functor C J_C.comp map a' d x'"
using A'.value_is_ide assms(3) by (unfold_locales, blast)
have d: "J_C.ide d" using J_C.ide_cod x.arrow by blast
let ?D = "J_C.Map d"
let ?\<chi> = "J_C.Map x"
let ?\<chi>' = "J_C.Map x'"
interpret D: diagram J C ?D
using ide_determines_diagram J_C.ide_cod x.arrow by blast
interpret \<chi>: cone J C ?D a ?\<chi>
using assms(1) d arrow_determines_cone by simp
interpret \<gamma>: constant_transformation J C g
using g \<chi>.ide_apex by (unfold_locales, auto)
interpret \<chi>og: vertical_composite J C A'.map \<chi>.A.map ?D \<gamma>.map ?\<chi>
using g C.comp_arr_dom C.comp_cod_arr \<gamma>.is_extensional by (unfold_locales, auto)
show ?thesis
proof
assume 0: "x.is_coext a' x' g"
show "?\<chi>' = D.cones_map g ?\<chi>"
proof -
have 1: "x' = x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g"
using 0 x.is_coext_def by blast
hence "?\<chi>' = J_C.Map x'"
using 0 x.is_coext_def by fast
moreover have "... = D.cones_map g ?\<chi>"
proof -
have "J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)) = x'"
proof -
have "J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)) =
x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g"
using d g cones_map_is_composition arrow_determines_cone(2) \<chi>.cone_axioms
x.arrow_from_functor_axioms
by auto
thus ?thesis by (metis 1)
qed
moreover have "J_C.arr (J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)))"
using 1 d g cones_map_is_composition preserves_arr arrow_determines_cone(2)
\<chi>.cone_axioms x.arrow_from_functor_axioms assms(3)
by auto
ultimately show ?thesis by auto
qed
ultimately show ?thesis by blast
qed
next
assume X': "?\<chi>' = D.cones_map g ?\<chi>"
show "x.is_coext a' x' g"
proof -
have 4: "J_C.seq x (map g)"
using g x.arrow mem_Collect_eq preserves_arr preserves_cod
by (elim C.in_homE, auto)
hence 1: "x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g =
J_C.MkArr (J_C.Dom (map g)) (J_C.Cod x)
(vertical_composite.map J C (J_C.Map (map g)) ?\<chi>)"
using J_C.comp_char [of x "map g"] by simp
have 2: "vertical_composite.map J C (J_C.Map (map g)) ?\<chi> = \<chi>og.map"
by (simp add: map_def \<gamma>.value_is_arr \<gamma>.natural_transformation_axioms)
have 3: "... = D.cones_map g ?\<chi>"
using g \<chi>og.map_simp_2 \<chi>.cone_axioms \<chi>og.is_extensional by auto
have "J_C.MkArr A'.map ?D ?\<chi>' = J_C.comp x (map g)"
proof -
have f1: "A'.map = J_C.Dom (map g)"
using \<gamma>.natural_transformation_axioms map_def g by auto
have "J_C.Map d = J_C.Cod x"
using x.arrow by auto
thus ?thesis using f1 X' 1 2 3 by argo
qed
moreover have "J_C.MkArr A'.map ?D ?\<chi>' = x'"
using d x' arrow_determines_cone by blast
ultimately show ?thesis
using g x.is_coext_def by simp
qed
qed
qed
end
locale right_adjoint_to_diagonal_functor =
C: category C +
J: category J +
J_C: functor_category J C +
\<Delta>: diagonal_functor J C +
"functor" J_C.comp C G +
Adj: meta_adjunction J_C.comp C \<Delta>.map G \<phi> \<psi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and G :: "('j, 'c) functor_category.arr \<Rightarrow> 'c"
and \<phi> :: "'c \<Rightarrow> ('j, 'c) functor_category.arr \<Rightarrow> 'c"
and \<psi> :: "('j, 'c) functor_category.arr \<Rightarrow> 'c \<Rightarrow> ('j, 'c) functor_category.arr" +
assumes adjoint: "adjoint_functors J_C.comp C \<Delta>.map G"
begin
interpretation S: replete_setcat .
interpretation Adj: adjunction J_C.comp C S.comp S.setp Adj.\<phi>C Adj.\<phi>D \<Delta>.map G
\<phi> \<psi> Adj.\<eta> Adj.\<epsilon> Adj.\<Phi> Adj.\<Psi>
using Adj.induces_adjunction by simp
text\<open>
A right adjoint @{term G} to a diagonal functor maps each object @{term d} of
\<open>[J, C]\<close> (corresponding to a diagram @{term D} of shape @{term J} in @{term C}
to an object of @{term C}. This object is the limit object, and the component at @{term d}
of the counit of the adjunction determines the limit cone.
\<close>
lemma gives_limit_cones:
assumes "diagram J C D"
shows "limit_cone J C D (G (J_C.MkIde D)) (J_C.Map (Adj.\<epsilon> (J_C.MkIde D)))"
proof -
interpret D: diagram J C D using assms by auto
let ?d = "J_C.MkIde D"
let ?a = "G ?d"
let ?x = "Adj.\<epsilon> ?d"
let ?\<chi> = "J_C.Map ?x"
have "diagram J C D" ..
hence 1: "J_C.ide ?d" using \<Delta>.diagram_determines_ide by auto
hence 2: "J_C.Map (J_C.MkIde D) = D"
using assms 1 J_C.in_homE \<Delta>.diagram_determines_ide(2) by simp
interpret x: terminal_arrow_from_functor C J_C.comp \<Delta>.map ?a ?d ?x
apply unfold_locales
apply (metis (no_types, lifting) "1" preserves_ide Adj.\<epsilon>_in_terms_of_\<psi>
Adj.\<epsilon>o_def Adj.\<epsilon>o_in_hom)
by (metis 1 Adj.has_terminal_arrows_from_functor(1)
terminal_arrow_from_functor.is_terminal)
have 3: "arrow_from_functor C J_C.comp \<Delta>.map ?a ?d ?x" ..
interpret \<chi>: cone J C D ?a ?\<chi>
using 1 2 3 \<Delta>.arrow_determines_cone [of ?d] by auto
have cone_\<chi>: "D.cone ?a ?\<chi>" ..
interpret \<chi>: limit_cone J C D ?a ?\<chi>
proof
fix a' \<chi>'
assume cone_\<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J C D a' \<chi>' using cone_\<chi>' by auto
let ?x' = "J_C.MkArr \<chi>'.A.map D \<chi>'"
interpret x': arrow_from_functor C J_C.comp \<Delta>.map a' ?d ?x'
using 1 2 by (metis \<Delta>.cone_determines_arrow(1) cone_\<chi>')
have "arrow_from_functor C J_C.comp \<Delta>.map a' ?d ?x'" ..
hence 4: "\<exists>!g. x.is_coext a' ?x' g"
using x.is_terminal by simp
have 5: "\<And>g. \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright> \<Longrightarrow> x.is_coext a' ?x' g \<longleftrightarrow> D.cones_map g ?\<chi> = \<chi>'"
using \<Delta>.coextension_iff_cones_map x'.arrow x.arrow_from_functor_axioms by auto
have 6: "\<And>g. x.is_coext a' ?x' g \<Longrightarrow> \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright>"
using x.is_coext_def by simp
show "\<exists>!g. \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright> \<and> D.cones_map g ?\<chi> = \<chi>'"
proof -
have "\<exists>g. \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright> \<and> D.cones_map g ?\<chi> = \<chi>'"
using 4 5 6 by meson
thus ?thesis
using 4 5 6 by blast
qed
qed
show "D.limit_cone ?a ?\<chi>" ..
qed
corollary gives_limits:
assumes "diagram J C D"
shows "diagram.has_as_limit J C D (G (J_C.MkIde D))"
using assms gives_limit_cones by fastforce
end
lemma (in category) has_limits_iff_left_adjoint_diagonal:
assumes "category J"
shows "has_limits_of_shape J \<longleftrightarrow>
left_adjoint_functor C (functor_category.comp J C) (diagonal_functor.map J C)"
proof -
interpret J: category J using assms by auto
interpret J_C: functor_category J C ..
interpret \<Delta>: diagonal_functor J C ..
show ?thesis
proof
assume A: "left_adjoint_functor C J_C.comp \<Delta>.map"
interpret \<Delta>: left_adjoint_functor C J_C.comp \<Delta>.map using A by auto
interpret Adj: meta_adjunction J_C.comp C \<Delta>.map \<Delta>.G \<Delta>.\<phi> \<Delta>.\<psi>
using \<Delta>.induces_meta_adjunction by auto
have 1: "adjoint_functors J_C.comp C \<Delta>.map \<Delta>.G"
using adjoint_functors_def \<Delta>.induces_meta_adjunction by blast
interpret G: right_adjoint_to_diagonal_functor J C \<Delta>.G \<Delta>.\<phi> \<Delta>.\<psi>
using 1 by unfold_locales auto
show "has_limits_of_shape J"
using A G.gives_limits has_limits_of_shape_def by blast
next
text\<open>
If @{term "has_limits J"}, then every diagram @{term D} from @{term J} to
@{term[source=true] C} has a limit cone.
This means that, for every object @{term d} of the functor category
\<open>[J, C]\<close>, there exists an object @{term a} of @{term C} and a terminal arrow from
\<open>\<Delta> a\<close> to @{term d} in \<open>[J, C]\<close>. The terminal arrow is given by the
limit cone.
\<close>
assume A: "has_limits_of_shape J"
show "left_adjoint_functor C J_C.comp \<Delta>.map"
proof
fix d
assume D: "J_C.ide d"
interpret D: diagram J C \<open>J_C.Map d\<close>
using D \<Delta>.ide_determines_diagram by auto
let ?D = "J_C.Map d"
have "diagram J C (J_C.Map d)" ..
from this obtain a \<chi> where limit: "limit_cone J C ?D a \<chi>"
using A has_limits_of_shape_def by blast
interpret A: constant_functor J C a
using limit by (simp add: Limit.cone_def limit_cone_def)
interpret \<chi>: limit_cone J C ?D a \<chi> using limit by simp
have cone_\<chi>: "cone J C ?D a \<chi>" ..
let ?x = "J_C.MkArr A.map ?D \<chi>"
interpret x: arrow_from_functor C J_C.comp \<Delta>.map a d ?x
using D cone_\<chi> \<Delta>.cone_determines_arrow by auto
have "terminal_arrow_from_functor C J_C.comp \<Delta>.map a d ?x"
proof
show "\<And>a' x'. arrow_from_functor C J_C.comp \<Delta>.map a' d x' \<Longrightarrow> \<exists>!g. x.is_coext a' x' g"
proof -
fix a' x'
assume x': "arrow_from_functor C J_C.comp \<Delta>.map a' d x'"
interpret x': arrow_from_functor C J_C.comp \<Delta>.map a' d x' using x' by auto
interpret A': constant_functor J C a'
by (unfold_locales, simp add: x'.arrow)
let ?\<chi>' = "J_C.Map x'"
interpret \<chi>': cone J C ?D a' ?\<chi>'
using D x' \<Delta>.arrow_determines_cone by auto
have cone_\<chi>': "cone J C ?D a' ?\<chi>'" ..
let ?g = "\<chi>.induced_arrow a' ?\<chi>'"
show "\<exists>!g. x.is_coext a' x' g"
proof
show "x.is_coext a' x' ?g"
proof (unfold x.is_coext_def)
have 1: "\<guillemotleft>?g : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map ?g \<chi> = ?\<chi>'"
using \<chi>.induced_arrow_def \<chi>.is_universal cone_\<chi>'
theI' [of "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = ?\<chi>'"]
by presburger
hence 2: "x' = ?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map ?g"
proof -
have "x' = J_C.MkArr A'.map ?D ?\<chi>'"
using D \<Delta>.arrow_determines_cone(2) x'.arrow_from_functor_axioms by auto
thus ?thesis
using 1 cone_\<chi> \<Delta>.cones_map_is_composition [of ?g a' a ?D \<chi>] by simp
qed
show "\<guillemotleft>?g : a' \<rightarrow> a\<guillemotright> \<and> x' = ?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map ?g"
using 1 2 by auto
qed
next
fix g
assume X: "x.is_coext a' x' g"
show "g = ?g"
proof -
have "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map g \<chi> = ?\<chi>'"
proof
show G: "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright>" using X x.is_coext_def by blast
show "D.cones_map g \<chi> = ?\<chi>'"
proof -
have "?\<chi>' = J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map g)"
using X x.is_coext_def [of a' x' g] by fast
also have "... = D.cones_map g \<chi>"
proof -
interpret map_g: constant_transformation J C g
using G by (unfold_locales, auto)
interpret \<chi>': vertical_composite J C
map_g.F.map A.map \<open>\<chi>.\<Phi>.Ya.Cop_S.Map d\<close>
map_g.map \<chi>
proof (intro_locales)
have "map_g.G.map = A.map"
using G by blast
thus "natural_transformation_axioms J (\<cdot>) map_g.F.map A.map map_g.map"
using map_g.natural_transformation_axioms
by (simp add: natural_transformation_def)
qed
have "J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map g) = vertical_composite.map J C map_g.map \<chi>"
proof -
have "J_C.seq ?x (\<Delta>.map g)"
using G x.arrow by auto
thus ?thesis
using G \<Delta>.map_def J_C.Map_comp' [of ?x "\<Delta>.map g"] by auto
qed
also have "... = D.cones_map g \<chi>"
using G cone_\<chi> \<chi>'.map_def map_g.map_def \<chi>.is_natural_2 \<chi>'.map_simp_2
by auto
finally show ?thesis by blast
qed
finally show ?thesis by auto
qed
qed
thus ?thesis
using cone_\<chi>' \<chi>.is_universal \<chi>.induced_arrow_def
theI_unique [of "\<lambda>g. \<guillemotleft>g : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map g \<chi> = ?\<chi>'" g]
by presburger
qed
qed
qed
qed
thus "\<exists>a x. terminal_arrow_from_functor C J_C.comp \<Delta>.map a d x" by auto
qed
qed
qed
section "Right Adjoint Functors Preserve Limits"
context right_adjoint_functor
begin
lemma preserves_limits:
fixes J :: "'j comp"
assumes "diagram J C E" and "diagram.has_as_limit J C E a"
shows "diagram.has_as_limit J D (G o E) (G a)"
proof -
text\<open>
From the assumption that @{term E} has a limit, obtain a limit cone @{term \<chi>}.
\<close>
interpret J: category J using assms(1) diagram_def by auto
interpret E: diagram J C E using assms(1) by auto
from assms(2) obtain \<chi> where \<chi>: "limit_cone J C E a \<chi>" by auto
interpret \<chi>: limit_cone J C E a \<chi> using \<chi> by auto
have a: "C.ide a" using \<chi>.ide_apex by auto
text\<open>
Form the @{term E}-image \<open>GE\<close> of the diagram @{term E}.
\<close>
interpret GE: composite_functor J C D E G ..
interpret GE: diagram J D GE.map ..
text\<open>Let \<open>G\<chi>\<close> be the @{term G}-image of the cone @{term \<chi>},
and note that it is a cone over \<open>GE\<close>.\<close>
let ?G\<chi> = "G o \<chi>"
interpret G\<chi>: cone J D GE.map \<open>G a\<close> ?G\<chi>
using \<chi>.cone_axioms preserves_cones by blast
text\<open>
Claim that \<open>G\<chi>\<close> is a limit cone for diagram \<open>GE\<close>.
\<close>
interpret G\<chi>: limit_cone J D GE.map \<open>G a\<close> ?G\<chi>
proof
text \<open>
Let @{term \<kappa>} be an arbitrary cone over \<open>GE\<close>.
\<close>
fix b \<kappa>
assume \<kappa>: "GE.cone b \<kappa>"
interpret \<kappa>: cone J D GE.map b \<kappa> using \<kappa> by auto
interpret Fb: constant_functor J C \<open>F b\<close>
apply unfold_locales
by (meson F_is_functor \<kappa>.ide_apex functor.preserves_ide)
interpret Adj: meta_adjunction C D F G \<phi> \<psi>
using induces_meta_adjunction by auto
interpret S: replete_setcat .
interpret Adj: adjunction C D S.comp S.setp
Adj.\<phi>C Adj.\<phi>D F G \<phi> \<psi> Adj.\<eta> Adj.\<epsilon> Adj.\<Phi> Adj.\<Psi>
using Adj.induces_adjunction by simp
text\<open>
For each arrow @{term j} of @{term J}, let @{term "\<chi>' j"} be defined to be
the adjunct of @{term "\<chi> j"}. We claim that @{term \<chi>'} is a cone over @{term E}.
\<close>
let ?\<chi>' = "\<lambda>j. if J.arr j then Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j) else C.null"
have cone_\<chi>': "E.cone (F b) ?\<chi>'"
proof
show "\<And>j. \<not>J.arr j \<Longrightarrow> ?\<chi>' j = C.null" by simp
fix j
assume j: "J.arr j"
show "C.dom (?\<chi>' j) = Fb.map (J.dom j)" using j \<psi>_in_hom by simp
show "C.cod (?\<chi>' j) = E (J.cod j)" using j \<psi>_in_hom by simp
show "E j \<cdot>\<^sub>C ?\<chi>' (J.dom j) = ?\<chi>' j"
proof -
have "E j \<cdot>\<^sub>C ?\<chi>' (J.dom j) = (E j \<cdot>\<^sub>C Adj.\<epsilon> (E (J.dom j))) \<cdot>\<^sub>C F (\<kappa> (J.dom j))"
using j C.comp_assoc by simp
also have "... = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> j)"
proof -
have "(E j \<cdot>\<^sub>C Adj.\<epsilon> (E (J.dom j))) \<cdot>\<^sub>C F (\<kappa> (J.dom j))
= (Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C Adj.FG.map (E j)) \<cdot>\<^sub>C F (\<kappa> (J.dom j))"
using j Adj.\<epsilon>.naturality [of "E j"] by fastforce
also have "... = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C Adj.FG.map (E j) \<cdot>\<^sub>C F (\<kappa> (J.dom j))"
using C.comp_assoc by simp
also have "... = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> j)"
proof -
have "Adj.FG.map (E j) \<cdot>\<^sub>C F (\<kappa> (J.dom j)) = F (GE.map j \<cdot>\<^sub>D \<kappa> (J.dom j))"
using j by simp
hence "Adj.FG.map (E j) \<cdot>\<^sub>C F (\<kappa> (J.dom j)) = F (\<kappa> j)"
using j \<kappa>.is_natural_1 by metis
thus ?thesis using j by simp
qed
finally show ?thesis by auto
qed
also have "... = ?\<chi>' j"
using j by simp
finally show ?thesis by auto
qed
show "?\<chi>' (J.cod j) \<cdot>\<^sub>C Fb.map j = ?\<chi>' j"
proof -
have "?\<chi>' (J.cod j) \<cdot>\<^sub>C Fb.map j = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> (J.cod j))"
using j Fb.value_is_ide Adj.\<epsilon>.preserves_hom C.comp_arr_dom [of "F (\<kappa> (J.cod j))"]
C.comp_assoc
by simp
also have "... = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> j)"
using j \<kappa>.is_natural_1 \<kappa>.is_natural_2 Adj.\<epsilon>.naturality J.arr_cod_iff_arr
by (metis J.cod_cod \<kappa>.A.map_simp)
also have "... = ?\<chi>' j" using j by simp
finally show ?thesis by auto
qed
qed
text\<open>
Using the universal property of the limit cone @{term \<chi>}, obtain the unique arrow
@{term f} that transforms @{term \<chi>} into @{term \<chi>'}.
\<close>
from this \<chi>.is_universal [of "F b" ?\<chi>'] obtain f
where f: "\<guillemotleft>f : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map f \<chi> = ?\<chi>'"
by auto
text\<open>
Let @{term g} be the adjunct of @{term f}, and show that @{term g} transforms
@{term G\<chi>} into @{term \<kappa>}.
\<close>
let ?g = "G f \<cdot>\<^sub>D Adj.\<eta> b"
have 1: "\<guillemotleft>?g : b \<rightarrow>\<^sub>D G a\<guillemotright>" using f \<kappa>.ide_apex by fastforce
moreover have "GE.cones_map ?g ?G\<chi> = \<kappa>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> GE.cones_map ?g ?G\<chi> j = \<kappa> j"
using 1 G\<chi>.cone_axioms \<kappa>.is_extensional by auto
moreover have "J.arr j \<Longrightarrow> GE.cones_map ?g ?G\<chi> j = \<kappa> j"
proof -
fix j
assume j: "J.arr j"
have "GE.cones_map ?g ?G\<chi> j = G (\<chi> j) \<cdot>\<^sub>D ?g"
using j 1 G\<chi>.cone_axioms mem_Collect_eq restrict_apply by auto
also have "... = G (\<chi> j \<cdot>\<^sub>C f) \<cdot>\<^sub>D Adj.\<eta> b"
using j f \<chi>.preserves_hom [of j "J.dom j" "J.cod j"] D.comp_assoc by fastforce
also have "... = G (E.cones_map f \<chi> j) \<cdot>\<^sub>D Adj.\<eta> b"
proof -
have "\<chi> j \<cdot>\<^sub>C f = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j)"
proof -
have "\<chi> j \<cdot>\<^sub>C f = E.cones_map f \<chi> j"
proof -
have "E.cone (C.cod f) \<chi>"
using f \<chi>.cone_axioms by blast
thus ?thesis
using \<chi>.is_extensional by simp
qed
also have "... = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j)"
using j f by simp
finally show ?thesis by blast
qed
thus ?thesis
using f mem_Collect_eq restrict_apply Adj.F.is_extensional by simp
qed
also have "... = (G (Adj.\<epsilon> (C.cod (E j))) \<cdot>\<^sub>D Adj.\<eta> (D.cod (GE.map j))) \<cdot>\<^sub>D \<kappa> j"
using j f Adj.\<eta>.naturality [of "\<kappa> j"] D.comp_assoc by auto
also have "... = D.cod (\<kappa> j) \<cdot>\<^sub>D \<kappa> j"
using j Adj.\<eta>\<epsilon>.triangle_G Adj.\<epsilon>_in_terms_of_\<psi> Adj.\<epsilon>o_def
Adj.\<eta>_in_terms_of_\<phi> Adj.\<eta>o_def Adj.unit_counit_G
by fastforce
also have "... = \<kappa> j"
using j D.comp_cod_arr by simp
finally show "GE.cones_map ?g ?G\<chi> j = \<kappa> j" by metis
qed
ultimately show "GE.cones_map ?g ?G\<chi> j = \<kappa> j" by auto
qed
ultimately have "\<guillemotleft>?g : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map ?g ?G\<chi> = \<kappa>" by auto
text\<open>
It remains to be shown that @{term g} is the unique such arrow.
Given any @{term g'} that transforms @{term G\<chi>} into @{term \<kappa>},
its adjunct transforms @{term \<chi>} into @{term \<chi>'}.
The adjunct of @{term g'} is therefore equal to @{term f},
which implies @{term g'} = @{term g}.
\<close>
moreover have "\<And>g'. \<guillemotleft>g' : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map g' ?G\<chi> = \<kappa> \<Longrightarrow> g' = ?g"
proof -
fix g'
assume g': "\<guillemotleft>g' : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map g' ?G\<chi> = \<kappa>"
have 1: "\<guillemotleft>\<psi> a g' : F b \<rightarrow>\<^sub>C a\<guillemotright>"
using g' a \<psi>_in_hom by simp
have 2: "E.cones_map (\<psi> a g') \<chi> = ?\<chi>'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j"
using 1 \<chi>.cone_axioms by auto
moreover have "J.arr j \<Longrightarrow> E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j"
proof -
fix j
assume j: "J.arr j"
have "E.cones_map (\<psi> a g') \<chi> j = \<chi> j \<cdot>\<^sub>C \<psi> a g'"
using 1 \<chi>.cone_axioms \<chi>.is_extensional by auto
also have "... = (\<chi> j \<cdot>\<^sub>C Adj.\<epsilon> a) \<cdot>\<^sub>C F g'"
using j a g' Adj.\<psi>_in_terms_of_\<epsilon> C.comp_assoc Adj.\<epsilon>_def by auto
also have "... = (Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (G (\<chi> j))) \<cdot>\<^sub>C F g'"
using j a g' Adj.\<epsilon>.naturality [of "\<chi> j"] by simp
also have "... = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j)"
using j a g' G\<chi>.cone_axioms C.comp_assoc by auto
finally show "E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j" by (simp add: j)
qed
ultimately show "E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j" by auto
qed
have "\<psi> a g' = f"
proof -
have "\<exists>!f. \<guillemotleft>f : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map f \<chi> = ?\<chi>'"
using cone_\<chi>' \<chi>.is_universal by simp
moreover have "\<guillemotleft>\<psi> a g' : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map (\<psi> a g') \<chi> = ?\<chi>'"
using 1 2 by simp
ultimately show ?thesis
using ex1E [of "\<lambda>f. \<guillemotleft>f : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map f \<chi> = ?\<chi>'" "\<psi> a g' = f"]
using 1 2 Adj.\<epsilon>.is_extensional C.comp_null(2) C.ex_un_null \<chi>.cone_axioms f
mem_Collect_eq restrict_apply
by blast
qed
hence "\<phi> b (\<psi> a g') = \<phi> b f" by auto
hence "g' = \<phi> b f" using \<chi>.ide_apex g' by (simp add: \<phi>_\<psi>)
moreover have "?g = \<phi> b f" using f Adj.\<phi>_in_terms_of_\<eta> \<kappa>.ide_apex Adj.\<eta>_def by auto
ultimately show "g' = ?g" by argo
qed
ultimately show "\<exists>!g. \<guillemotleft>g : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map g ?G\<chi> = \<kappa>" by blast
qed
have "GE.limit_cone (G a) ?G\<chi>" ..
thus ?thesis by auto
qed
end
section "Special Kinds of Limits"
subsection "Terminal Objects"
text\<open>
An object of a category @{term C} is a terminal object if and only if it is a limit of the
empty diagram in @{term C}.
\<close>
locale empty_diagram =
diagram J C D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c" +
assumes is_empty: "\<not>J.arr j"
begin
lemma has_as_limit_iff_terminal:
shows "has_as_limit a \<longleftrightarrow> C.terminal a"
proof
assume a: "has_as_limit a"
show "C.terminal a"
proof
have "\<exists>\<chi>. limit_cone a \<chi>" using a by auto
from this obtain \<chi> where \<chi>: "limit_cone a \<chi>" by blast
interpret \<chi>: limit_cone J C D a \<chi> using \<chi> by auto
have cone_\<chi>: "cone a \<chi>" ..
show "C.ide a" using \<chi>.ide_apex by auto
have 1: "\<chi> = (\<lambda>j. C.null)" using is_empty \<chi>.is_extensional by auto
show "\<And>a'. C.ide a' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>"
proof -
fix a'
assume a': "C.ide a'"
interpret A': constant_functor J C a'
apply unfold_locales using a' by auto
let ?\<chi>' = "\<lambda>j. C.null"
have cone_\<chi>': "cone a' ?\<chi>'"
using a' is_empty apply unfold_locales by auto
hence "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> cones_map f \<chi> = ?\<chi>'"
using \<chi>.is_universal by force
moreover have "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> cones_map f \<chi> = ?\<chi>'"
using 1 cone_\<chi> by auto
ultimately show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" by blast
qed
qed
next
assume a: "C.terminal a"
show "has_as_limit a"
proof -
let ?\<chi> = "\<lambda>j. C.null"
have "C.ide a" using a C.terminal_def by simp
interpret A: constant_functor J C a
apply unfold_locales using \<open>C.ide a\<close> by simp
interpret \<chi>: cone J C D a ?\<chi>
using \<open>C.ide a\<close> is_empty by (unfold_locales, auto)
have cone_\<chi>: "cone a ?\<chi>" ..
have 1: "\<And>a' \<chi>'. cone a' \<chi>' \<Longrightarrow> \<chi>' = (\<lambda>j. C.null)"
proof -
fix a' \<chi>'
assume \<chi>': "cone a' \<chi>'"
interpret \<chi>': cone J C D a' \<chi>' using \<chi>' by auto
show "\<chi>' = (\<lambda>j. C.null)"
using is_empty \<chi>'.is_extensional by metis
qed
have "limit_cone a ?\<chi>"
proof
fix a' \<chi>'
assume \<chi>': "cone a' \<chi>'"
have 2: "\<chi>' = (\<lambda>j. C.null)" using 1 \<chi>' by simp
interpret \<chi>': cone J C D a' \<chi>' using \<chi>' by auto
have "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" using a C.terminal_def \<chi>'.ide_apex by simp
moreover have "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> cones_map f ?\<chi> = \<chi>'"
using 1 2 cones_map_mapsto cone_\<chi> \<chi>'.cone_axioms mem_Collect_eq by blast
ultimately show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> cones_map f (\<lambda>j. C.null) = \<chi>'"
by blast
qed
thus ?thesis by auto
qed
qed
end
subsection "Products"
text\<open>
A \emph{product} in a category @{term C} is a limit of a discrete diagram in @{term C}.
\<close>
locale discrete_diagram =
J: category J +
diagram J C D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c" +
assumes is_discrete: "J.arr = J.ide"
begin
abbreviation mkCone
where "mkCone F \<equiv> (\<lambda>j. if J.arr j then F j else C.null)"
lemma cone_mkCone:
assumes "C.ide a" and "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>F j : a \<rightarrow> D j\<guillemotright>"
shows "cone a (mkCone F)"
proof -
interpret A: constant_functor J C a
using assms(1) by unfold_locales auto
show "cone a (mkCone F)"
using assms(2) is_discrete
apply unfold_locales
apply auto
apply (metis C.in_homE C.comp_cod_arr)
using C.comp_arr_ide by fastforce
qed
lemma mkCone_cone:
assumes "cone a \<pi>"
shows "mkCone \<pi> = \<pi>"
proof -
interpret \<pi>: cone J C D a \<pi>
using assms by auto
show "mkCone \<pi> = \<pi>" using \<pi>.is_extensional by auto
qed
end
text\<open>
The following locale defines a discrete diagram in a category @{term C},
given an index set @{term I} and a function @{term D} mapping @{term I}
to objects of @{term C}. Here we obtain the diagram shape @{term J}
using a discrete category construction that allows us to directly identify
the objects of @{term J} with the elements of @{term I}, however this construction
can only be applied in case the set @{term I} is not the universe of its
element type.
\<close>
locale discrete_diagram_from_map =
J: discrete_category I null +
C: category C
for I :: "'i set"
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'i \<Rightarrow> 'c"
and null :: 'i +
assumes maps_to_ide: "i \<in> I \<Longrightarrow> C.ide (D i)"
begin
definition map
where "map j \<equiv> if J.arr j then D j else C.null"
end
sublocale discrete_diagram_from_map \<subseteq> discrete_diagram J.comp C map
using map_def maps_to_ide J.arr_char J.Null_not_in_Obj J.null_char
by unfold_locales auto
locale product_cone =
J: category J +
C: category C +
D: discrete_diagram J C D +
limit_cone J C D a \<pi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and a :: 'c
and \<pi> :: "'j \<Rightarrow> 'c"
begin
lemma is_cone:
shows "D.cone a \<pi>" ..
text\<open>
The following versions of @{prop is_universal} and @{prop induced_arrowI}
from the \<open>limit_cone\<close> locale are specialized to the case in which the
underlying diagram is a product diagram.
\<close>
lemma is_universal':
assumes "C.ide b" and "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>F j: b \<rightarrow> D j\<guillemotright>"
shows "\<exists>!f. \<guillemotleft>f : b \<rightarrow> a\<guillemotright> \<and> (\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j)"
proof -
let ?\<chi> = "D.mkCone F"
interpret B: constant_functor J C b
using assms(1) by unfold_locales auto
have cone_\<chi>: "D.cone b ?\<chi>"
using assms D.is_discrete D.cone_mkCone by blast
interpret \<chi>: cone J C D b ?\<chi> using cone_\<chi> by auto
have "\<exists>!f. \<guillemotleft>f : b \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<pi> = ?\<chi>"
using cone_\<chi> is_universal by force
moreover have
"\<And>f. \<guillemotleft>f : b \<rightarrow> a\<guillemotright> \<Longrightarrow> D.cones_map f \<pi> = ?\<chi> \<longleftrightarrow> (\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j)"
proof -
fix f
assume f: "\<guillemotleft>f : b \<rightarrow> a\<guillemotright>"
show "D.cones_map f \<pi> = ?\<chi> \<longleftrightarrow> (\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j)"
proof
assume 1: "D.cones_map f \<pi> = ?\<chi>"
show "\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j"
proof -
have "\<And>j. J.arr j \<Longrightarrow> \<pi> j \<cdot> f = F j"
proof -
fix j
assume j: "J.arr j"
have "\<pi> j \<cdot> f = D.cones_map f \<pi> j"
using j f cone_axioms by force
also have "... = F j" using j 1 by simp
finally show "\<pi> j \<cdot> f = F j" by auto
qed
thus ?thesis by auto
qed
next
assume 1: "\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j"
show "D.cones_map f \<pi> = ?\<chi>"
using 1 f is_cone \<chi>.is_extensional D.is_discrete is_cone cone_\<chi> by auto
qed
qed
ultimately show ?thesis by blast
qed
abbreviation induced_arrow' :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) \<Rightarrow> 'c"
where "induced_arrow' b F \<equiv> induced_arrow b (D.mkCone F)"
lemma induced_arrowI':
assumes "C.ide b" and "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>F j : b \<rightarrow> D j\<guillemotright>"
shows "\<And>j. J.arr j \<Longrightarrow> \<pi> j \<cdot> induced_arrow' b F = F j"
proof -
interpret B: constant_functor J C b
using assms(1) by unfold_locales auto
interpret \<chi>: cone J C D b \<open>D.mkCone F\<close>
using assms D.cone_mkCone by blast
have cone_\<chi>: "D.cone b (D.mkCone F)" ..
hence 1: "D.cones_map (induced_arrow' b F) \<pi> = D.mkCone F"
using induced_arrowI by blast
fix j
assume j: "J.arr j"
have "\<pi> j \<cdot> induced_arrow' b F = D.cones_map (induced_arrow' b F) \<pi> j"
using induced_arrowI(1) cone_\<chi> is_cone is_extensional by force
also have "... = F j"
using j 1 by auto
finally show "\<pi> j \<cdot> induced_arrow' b F = F j"
by auto
qed
end
context discrete_diagram
begin
lemma product_coneI:
assumes "limit_cone a \<pi>"
shows "product_cone J C D a \<pi>"
by (meson assms discrete_diagram_axioms functor_axioms functor_def product_cone.intro)
end
context category
begin
definition has_as_product
where "has_as_product J D a \<equiv> (\<exists>\<pi>. product_cone J C D a \<pi>)"
lemma product_is_ide:
assumes "has_as_product J D a"
shows "ide a"
proof -
obtain \<pi> where \<pi>: "product_cone J C D a \<pi>"
using assms has_as_product_def by blast
interpret \<pi>: product_cone J C D a \<pi>
using \<pi> by auto
show ?thesis using \<pi>.ide_apex by auto
qed
text\<open>
A category has @{term I}-indexed products for an @{typ 'i}-set @{term I}
if every @{term I}-indexed discrete diagram has a product.
In order to reap the benefits of being able to directly identify the elements
of a set I with the objects of discrete category it generates (thereby avoiding
the use of coercion maps), it is necessary to assume that @{term "I \<noteq> UNIV"}.
If we want to assert that a category has products indexed by the universe of
some type @{typ 'i}, we have to pass to a larger type, such as @{typ "'i option"}.
\<close>
definition has_products
where "has_products (I :: 'i set) \<equiv>
I \<noteq> UNIV \<and>
(\<forall>J D. discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I
\<longrightarrow> (\<exists>a. has_as_product J D a))"
lemma ex_productE:
assumes "\<exists>a. has_as_product J D a"
obtains a \<pi> where "product_cone J C D a \<pi>"
using assms has_as_product_def someI_ex [of "\<lambda>a. has_as_product J D a"] by metis
lemma has_products_if_has_limits:
assumes "has_limits (undefined :: 'j)" and "I \<noteq> (UNIV :: 'j set)"
shows "has_products I"
proof (unfold has_products_def, intro conjI allI impI)
show "I \<noteq> UNIV" by fact
fix J D
assume D: "discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I"
interpret D: discrete_diagram J C D
using D by simp
have 1: "\<exists>a. D.has_as_limit a"
using assms D D.diagram_axioms D.J.category_axioms
by (simp add: has_limits_of_shape_def has_limits_def)
show "\<exists>a. has_as_product J D a"
using 1 has_as_product_def D.product_coneI by blast
qed
lemma has_finite_products_if_has_finite_limits:
assumes "\<And>J :: 'j comp. (finite (Collect (partial_magma.arr J))) \<Longrightarrow> has_limits_of_shape J"
and "finite (I :: 'j set)" and "I \<noteq> UNIV"
shows "has_products I"
proof (unfold has_products_def, intro conjI allI impI)
show "I \<noteq> UNIV" by fact
fix J D
assume D: "discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I"
interpret D: discrete_diagram J C D
using D by simp
have 1: "\<exists>a. D.has_as_limit a"
using assms D has_limits_of_shape_def D.diagram_axioms by auto
show "\<exists>a. has_as_product J D a"
using 1 has_as_product_def D.product_coneI by blast
qed
lemma has_products_preserved_by_bijection:
assumes "has_products I" and "bij_betw \<phi> I I'" and "I' \<noteq> UNIV"
shows "has_products I'"
proof (unfold has_products_def, intro conjI allI impI)
show "I' \<noteq> UNIV" by fact
show "\<And>J' D'. discrete_diagram J' C D' \<and> Collect (partial_magma.arr J') = I'
\<Longrightarrow> \<exists>a. has_as_product J' D' a"
proof -
fix J' D'
assume 1: "discrete_diagram J' C D' \<and> Collect (partial_magma.arr J') = I'"
interpret J': category J'
using 1 by (simp add: discrete_diagram_def)
interpret D': discrete_diagram J' C D'
using 1 by simp
interpret J: discrete_category I \<open>SOME x. x \<notin> I\<close>
using assms has_products_def [of I] someI_ex [of "\<lambda>x. x \<notin> I"]
by unfold_locales auto
have 2: "Collect J.arr = I \<and> Collect J'.arr = I'"
using 1 by auto
have \<phi>: "bij_betw \<phi> (Collect J.arr) (Collect J'.arr)"
using 2 assms(2) by simp
let ?\<phi> = "\<lambda>j. if J.arr j then \<phi> j else J'.null"
let ?\<phi>' = "\<lambda>j'. if J'.arr j' then the_inv_into I \<phi> j' else J.null"
interpret \<phi>: "functor" J.comp J' ?\<phi>
proof -
have "\<phi> ` I = I'"
using \<phi> 2 bij_betw_def [of \<phi> I I'] by simp
hence "\<And>j. J.arr j \<Longrightarrow> J'.arr (?\<phi> j)"
using 1 D'.is_discrete by auto
thus "functor J.comp J' ?\<phi>"
using D'.is_discrete J.is_discrete J.seqE
by unfold_locales auto
qed
interpret \<phi>': "functor" J' J.comp ?\<phi>'
proof -
have "the_inv_into I \<phi> ` I' = I"
using assms(2) \<phi> bij_betw_the_inv_into bij_betw_imp_surj_on by metis
hence "\<And>j'. J'.arr j' \<Longrightarrow> J.arr (?\<phi>' j')"
using 2 D'.is_discrete J.is_discrete by auto
thus "functor J' J.comp ?\<phi>'"
using D'.is_discrete J.is_discrete J'.seqE
by unfold_locales auto
qed
let ?D = "\<lambda>i. D' (\<phi> i)"
interpret D: discrete_diagram_from_map I C ?D \<open>SOME j. j \<notin> I\<close>
using assms 1 D'.is_discrete bij_betw_imp_surj_on \<phi>.preserves_ide
by unfold_locales auto
obtain a where a: "has_as_product J.comp D.map a"
using assms D.discrete_diagram_axioms has_products_def [of I] by auto
obtain \<pi> where \<pi>: "product_cone J.comp C D.map a \<pi>"
using a has_as_product_def by blast
interpret \<pi>: product_cone J.comp C D.map a \<pi>
using \<pi> by simp
let ?\<pi>' = "\<pi> o ?\<phi>'"
interpret A: constant_functor J' C a
using \<pi>.ide_apex by unfold_locales simp
interpret \<pi>': natural_transformation J' C A.map D' ?\<pi>'
proof -
have "\<pi>.A.map \<circ> ?\<phi>' = A.map"
using \<phi> A.map_def \<phi>'.preserves_arr \<pi>.A.is_extensional J.not_arr_null by auto
moreover have "D.map \<circ> ?\<phi>' = D'"
proof
fix j'
have "J'.arr j' \<Longrightarrow> (D.map \<circ> ?\<phi>') j' = D' j'"
proof -
assume 2: "J'.arr j'"
have 3: "inj_on \<phi> I"
using assms(2) bij_betw_imp_inj_on by auto
have "\<phi> ` I = I'"
by (metis (no_types) assms(2) bij_betw_imp_surj_on)
hence "\<phi> ` I = Collect J'.arr"
using 1 by force
thus ?thesis
using 2 3 D.map_def \<phi>'.preserves_arr f_the_inv_into_f by fastforce
qed
moreover have "\<not> J'.arr j' \<Longrightarrow> (D.map \<circ> ?\<phi>') j' = D' j'"
using D.is_extensional D'.is_extensional
by (simp add: J.Null_not_in_Obj J.null_char)
ultimately show "(D.map \<circ> ?\<phi>') j' = D' j'" by blast
qed
ultimately show "natural_transformation J' C A.map D' ?\<pi>'"
using \<pi>.natural_transformation_axioms \<phi>'.as_nat_trans.natural_transformation_axioms
horizontal_composite [of J' J.comp ?\<phi>' ?\<phi>' ?\<phi>' C \<pi>.A.map D.map \<pi>]
by simp
qed
interpret \<pi>': cone J' C D' a ?\<pi>' ..
interpret \<pi>': product_cone J' C D' a ?\<pi>'
proof
fix a' \<chi>'
assume \<chi>': "D'.cone a' \<chi>'"
interpret \<chi>': cone J' C D' a' \<chi>'
using \<chi>' by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f (\<pi> \<circ> ?\<phi>') = \<chi>'"
proof -
let ?\<chi> = "\<chi>' o ?\<phi>"
interpret A': constant_functor J.comp C a'
using \<chi>'.ide_apex by unfold_locales simp
interpret \<chi>: natural_transformation J.comp C A'.map D.map ?\<chi>
proof -
have "\<chi>'.A.map \<circ> ?\<phi> = A'.map"
using \<phi> \<phi>.preserves_arr A'.map_def \<chi>'.A.is_extensional by auto
moreover have "D' \<circ> ?\<phi> = D.map"
using \<phi> D.map_def D'.is_extensional by auto
ultimately show "natural_transformation J.comp C A'.map D.map ?\<chi>"
using \<chi>'.natural_transformation_axioms
\<phi>.as_nat_trans.natural_transformation_axioms
horizontal_composite [of J.comp J' ?\<phi> ?\<phi> ?\<phi> C \<chi>'.A.map D' \<chi>']
by simp
qed
interpret \<chi>: cone J.comp C D.map a' ?\<chi> ..
have *: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<pi> = ?\<chi>"
using \<pi>.is_universal \<chi>.cone_axioms by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f ?\<pi>' = \<chi>'"
proof -
have "\<exists>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f ?\<pi>' = \<chi>'"
proof -
obtain f where f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<pi> = ?\<chi>"
using * by blast
have "D'.cones_map f ?\<pi>' = \<chi>'"
proof
fix j'
show "D'.cones_map f ?\<pi>' j' = \<chi>' j'"
proof (cases "J'.arr j'")
assume j': "\<not> J'.arr j'"
show "D'.cones_map f ?\<pi>' j' = \<chi>' j'"
using f j' \<chi>'.is_extensional \<pi>'.cone_axioms by auto
next
assume j': "J'.arr j'"
show "D'.cones_map f ?\<pi>' j' = \<chi>' j'"
proof -
have "D'.cones_map f ?\<pi>' j' = \<pi> (the_inv_into I \<phi> j') \<cdot> f"
using f j' \<pi>'.cone_axioms by auto
also have "... = D.cones_map f \<pi> (the_inv_into I \<phi> j')"
proof -
have "arr f \<and> dom f = a' \<and> cod f = a"
using f by blast
thus ?thesis
using \<phi>'.preserves_arr \<pi>.cone_\<chi> j' by auto
qed
also have "... = (\<chi>' \<circ> ?\<phi>) (the_inv_into I \<phi> j')"
using f by simp
also have "... = \<chi>' j'"
using assms(2) j' 2 bij_betw_def [of \<phi> I I'] bij_betw_imp_inj_on
\<phi>'.preserves_arr f_the_inv_into_f
by fastforce
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using f by blast
qed
moreover have "\<And>f f'. \<lbrakk> \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>; D'.cones_map f ?\<pi>' = \<chi>';
\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>; D'.cones_map f' ?\<pi>' = \<chi>' \<rbrakk>
\<Longrightarrow> f = f'"
proof -
fix f f'
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" and f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>"
and f\<chi>': "D'.cones_map f ?\<pi>' = \<chi>'" and f'\<chi>': "D'.cones_map f' ?\<pi>' = \<chi>'"
have "D.cones_map f \<pi> = \<chi>' \<circ> ?\<phi> \<and> D.cones_map f' \<pi> = \<chi>' o ?\<phi>"
proof (intro conjI)
show "D.cones_map f \<pi> = \<chi>' \<circ> ?\<phi>"
proof
fix j
have "\<not> J.arr j \<Longrightarrow> D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
using f f\<chi>' \<pi>.cone_axioms \<chi>.is_extensional by auto
moreover have "J.arr j \<Longrightarrow> D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
proof -
assume j: "J.arr j"
have 1: "j = the_inv_into I \<phi> (\<phi> j)"
using assms(2) j \<phi> the_inv_into_f_f bij_betw_imp_inj_on J.arr_char
by metis
have "D.cones_map f \<pi> j = D.cones_map f \<pi> (the_inv_into I \<phi> (\<phi> j))"
using 1 by simp
also have "... = (\<chi>' \<circ> ?\<phi>) j"
using f j f\<chi>' 1 \<pi>.cone_axioms \<pi>'.cone_axioms \<phi>.preserves_arr by auto
finally show "D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
ultimately show "D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
show "D.cones_map f' \<pi> = \<chi>' \<circ> ?\<phi>"
proof
fix j
have "\<not> J.arr j \<Longrightarrow> D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
using f' f\<chi>' \<pi>.cone_axioms \<chi>.is_extensional by auto
moreover have "J.arr j \<Longrightarrow> D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
proof -
assume j: "J.arr j"
have 1: "j = the_inv_into I \<phi> (\<phi> j)"
using assms(2) j \<phi> the_inv_into_f_f bij_betw_imp_inj_on J.arr_char
by metis
have "D.cones_map f' \<pi> j = D.cones_map f' \<pi> (the_inv_into I \<phi> (\<phi> j))"
using 1 by simp
also have "... = (\<chi>' \<circ> ?\<phi>) j"
using f' j f'\<chi>' 1 \<pi>.cone_axioms \<pi>'.cone_axioms \<phi>.preserves_arr by auto
finally show "D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
ultimately show "D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
qed
thus "f = f'"
using f f' * by auto
qed
ultimately show ?thesis by blast
qed
qed
qed
have "has_as_product J' D' a"
using has_as_product_def \<pi>'.product_cone_axioms by auto
thus "\<exists>a. has_as_product J' D' a" by blast
qed
qed
lemma ide_is_unary_product:
assumes "ide a"
shows "\<And>m n :: nat. m \<noteq> n \<Longrightarrow> has_as_product (discrete_category.comp {m :: nat} (n :: nat))
(\<lambda>i. if i = m then a else null) a"
proof -
fix m n :: nat
assume neq: "m \<noteq> n"
have "{m :: nat} \<noteq> UNIV"
proof -
have "finite {m :: nat}" by simp
moreover have "\<not>finite (UNIV :: nat set)" by simp
ultimately show ?thesis by fastforce
qed
interpret J: discrete_category "{m :: nat}" \<open>n :: nat\<close>
using neq \<open>{m :: nat} \<noteq> UNIV\<close> by unfold_locales auto
let ?D = "\<lambda>i. if i = m then a else null"
interpret D: discrete_diagram J.comp C ?D
apply unfold_locales
using assms J.null_char neq
apply auto
by metis
interpret A: constant_functor J.comp C a
using assms by unfold_locales auto
show "has_as_product J.comp ?D a"
proof (unfold has_as_product_def)
let ?\<pi> = "\<lambda>i :: nat. if i = m then a else null"
interpret \<pi>: natural_transformation J.comp C A.map ?D ?\<pi>
using assms J.arr_char J.dom_char J.cod_char
by unfold_locales auto
interpret \<pi>: cone J.comp C ?D a ?\<pi> ..
interpret \<pi>: product_cone J.comp C ?D a ?\<pi>
proof
fix a' \<chi>'
assume \<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J.comp C ?D a' \<chi>' using \<chi>' by auto
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'"
proof
show "\<guillemotleft>\<chi>' m : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map (\<chi>' m) ?\<pi> = \<chi>'"
proof
show 1: "\<guillemotleft>\<chi>' m : a' \<rightarrow> a\<guillemotright>"
using \<chi>'.preserves_hom \<chi>'.A.map_def J.arr_char J.dom_char J.cod_char
by auto
show "D.cones_map (\<chi>' m) ?\<pi> = \<chi>'"
proof
fix j
show "D.cones_map (\<chi>' m) (\<lambda>i. if i = m then a else null) j = \<chi>' j"
using J.arr_char J.dom_char J.cod_char J.ide_char \<pi>.cone_axioms comp_cod_arr
apply (cases "j = m")
apply simp
using \<chi>'.is_extensional by simp
qed
qed
show "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>' \<Longrightarrow> f = \<chi>' m"
proof -
fix f
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'"
show "f = \<chi>' m"
using assms f \<chi>'.preserves_hom J.arr_char J.dom_char J.cod_char \<pi>.cone_axioms
comp_cod_arr
by auto
qed
qed
qed
show "\<exists>\<pi>. product_cone J.comp C (\<lambda>i. if i = m then a else null) a \<pi>"
using \<pi>.product_cone_axioms by blast
qed
qed
lemma has_unary_products:
assumes "card I = 1" and "I \<noteq> UNIV"
shows "has_products I"
proof -
obtain i where i: "I = {i}"
using assms card_1_singletonE by blast
obtain n where n: "n \<notin> I"
using assms by auto
have "has_products {1 :: nat}"
proof (unfold has_products_def, intro conjI)
show "{1 :: nat} \<noteq> UNIV" by auto
show "\<forall>(J :: nat comp) D.
discrete_diagram J (\<cdot>) D \<and> Collect (partial_magma.arr J) = {1}
\<longrightarrow> (\<exists>a. has_as_product J D a)"
proof
fix J :: "nat comp"
show "\<forall>D. discrete_diagram J (\<cdot>) D \<and> Collect (partial_magma.arr J) = {1}
\<longrightarrow> (\<exists>a. has_as_product J D a)"
proof (intro allI impI)
fix D
assume D: "discrete_diagram J (\<cdot>) D \<and> Collect (partial_magma.arr J) = {1}"
interpret D: discrete_diagram J C D
using D by auto
interpret J: discrete_category \<open>{1 :: nat}\<close> D.J.null
by (metis D D.J.not_arr_null discrete_category.intro mem_Collect_eq)
have 1: "has_as_product J.comp D (D 1)"
proof -
have "has_as_product J.comp (\<lambda>i. if i = 1 then D 1 else null) (D 1)"
using ide_is_unary_product D.preserves_ide D.is_discrete D J.null_char
by (metis J.Null_not_in_Obj insertCI mem_Collect_eq)
moreover have "D = (\<lambda>i. if i = 1 then D 1 else null)"
proof
fix j
show "D j = (if j = 1 then D 1 else null)"
using D D.is_extensional by auto
qed
ultimately show ?thesis by simp
qed
moreover have "J = J.comp"
proof -
have "\<And>j j'. J j j' = J.comp j j'"
proof -
fix j j'
show "J j j' = J.comp j j'"
using J.comp_char D.is_discrete D
by (metis D.J.category_axioms D.J.ext D.J.ide_def J.null_char
category.seqE mem_Collect_eq)
qed
thus ?thesis by auto
qed
ultimately show "\<exists>a. has_as_product J D a" by auto
qed
qed
qed
moreover have "bij_betw (\<lambda>k. if k = 1 then i else n) {1 :: nat} I"
using i by (intro bij_betwI') auto
ultimately show "has_products I"
using assms has_products_preserved_by_bijection by blast
qed
end
subsection "Equalizers"
text\<open>
An \emph{equalizer} in a category @{term C} is a limit of a parallel pair
of arrows in @{term C}.
\<close>
locale parallel_pair_diagram =
J: parallel_pair +
C: category C
for C :: "'c comp" (infixr "\<cdot>" 55)
and f0 :: 'c
and f1 :: 'c +
assumes is_parallel: "C.par f0 f1"
begin
no_notation J.comp (infixr "\<cdot>" 55)
notation J.comp (infixr "\<cdot>\<^sub>J" 55)
definition map
where "map \<equiv> (\<lambda>j. if j = J.Zero then C.dom f0
else if j = J.One then C.cod f0
else if j = J.j0 then f0
else if j = J.j1 then f1
else C.null)"
lemma map_simp:
shows "map J.Zero = C.dom f0"
and "map J.One = C.cod f0"
and "map J.j0 = f0"
and "map J.j1 = f1"
proof -
show "map J.Zero = C.dom f0"
using map_def by metis
show "map J.One = C.cod f0"
using map_def J.Zero_not_eq_One by metis
show "map J.j0 = f0"
using map_def J.Zero_not_eq_j0 J.One_not_eq_j0 by metis
show "map J.j1 = f1"
using map_def J.Zero_not_eq_j1 J.One_not_eq_j1 J.j0_not_eq_j1 by metis
qed
end
sublocale parallel_pair_diagram \<subseteq> diagram J.comp C map
apply unfold_locales
apply (simp add: J.arr_char map_def)
using map_def is_parallel J.arr_char J.cod_simp J.dom_simp
apply auto[2]
proof -
show 1: "\<And>j. J.arr j \<Longrightarrow> C.cod (map j) = map (J.cod j)"
using is_parallel map_simp(1-4) J.arr_char by auto
next
fix j j'
assume jj': "J.seq j' j"
show "map (j' \<cdot>\<^sub>J j) = map j' \<cdot> map j"
proof -
have 1: "(j = J.Zero \<and> j' \<noteq> J.One) \<or> (j \<noteq> J.Zero \<and> j' = J.One)"
using jj' J.seq_char by blast
moreover have "j = J.Zero \<and> j' \<noteq> J.One \<Longrightarrow> ?thesis"
by (metis (no_types, lifting) C.arr_dom_iff_arr C.comp_arr_dom C.dom_dom
J.comp_simp(1) jj' map_simp(1,3-4) J.arr_char is_parallel)
moreover have "j \<noteq> J.Zero \<and> j' = J.One \<Longrightarrow> ?thesis"
by (metis (no_types, lifting) C.comp_arr_dom C.comp_cod_arr C.seqE J.comp_char jj'
map_simp(2-4) J.arr_char is_parallel)
ultimately show ?thesis by blast
qed
qed
context parallel_pair_diagram
begin
definition mkCone
where "mkCone e \<equiv> \<lambda>j. if J.arr j then if j = J.Zero then e else f0 \<cdot> e else C.null"
abbreviation is_equalized_by
where "is_equalized_by e \<equiv> C.seq f0 e \<and> f0 \<cdot> e = f1 \<cdot> e"
abbreviation has_as_equalizer
where "has_as_equalizer e \<equiv> limit_cone (C.dom e) (mkCone e)"
lemma cone_mkCone:
assumes "is_equalized_by e"
shows "cone (C.dom e) (mkCone e)"
proof -
interpret E: constant_functor J.comp C \<open>C.dom e\<close>
apply unfold_locales using assms by auto
show "cone (C.dom e) (mkCone e)"
using assms mkCone_def apply unfold_locales
apply auto[2]
using C.dom_comp C.seqE C.cod_comp J.Zero_not_eq_One J.arr_char' J.cod_char map_def
apply (metis (no_types, lifting) C.not_arr_null parallel_pair.cod_simp(1) preserves_arr)
proof -
show "\<And>j. J.arr j \<Longrightarrow> map j \<cdot> mkCone e (J.dom j) = mkCone e j"
proof -
fix j
assume j: "J.arr j"
have 1: "j = J.Zero \<or> map j \<cdot> mkCone e (J.dom j) = mkCone e j"
using assms j mkCone_def C.cod_comp
by (metis (no_types, lifting) C.comp_cod_arr J.arr_char J.dom_char map_def
J.dom_simp(2))
show "map j \<cdot> mkCone e (J.dom j) = mkCone e j"
by (metis (no_types, lifting) 1 C.comp_cod_arr C.seqE assms j map_simp(1)
mkCone_def J.dom_simp(1))
qed
next
show "\<And>j. J.arr j \<Longrightarrow> mkCone e (J.cod j) \<cdot> E.map j = mkCone e j"
by (metis (no_types, lifting) C.comp_arr_dom C.comp_assoc C.seqE
J.arr_cod_iff_arr J.seqI J.seq_char assms E.map_simp mkCone_def
J.cod_simp(1) J.dom_simp(1))
qed
qed
lemma is_equalized_by_cone:
assumes "cone a \<chi>"
shows "is_equalized_by (\<chi> (J.Zero))"
proof -
interpret \<chi>: cone J.comp C map a \<chi>
using assms by auto
show ?thesis
by (metis (no_types, lifting) J.arr_char J.cod_char cone_def
\<chi>.component_in_hom \<chi>.is_natural_1 \<chi>.naturality assms C.in_homE
constant_functor.map_simp J.dom_simp(3-4) map_simp(3-4))
qed
lemma mkCone_cone:
assumes "cone a \<chi>"
shows "mkCone (\<chi> J.Zero) = \<chi>"
proof -
interpret \<chi>: cone J.comp C map a \<chi>
using assms by auto
have 1: "is_equalized_by (\<chi> J.Zero)"
using assms is_equalized_by_cone by blast
show ?thesis
proof
fix j
have "j = J.Zero \<Longrightarrow> mkCone (\<chi> J.Zero) j = \<chi> j"
using mkCone_def \<chi>.is_extensional by simp
moreover have "j = J.One \<or> j = J.j0 \<or> j = J.j1 \<Longrightarrow> mkCone (\<chi> J.Zero) j = \<chi> j"
using J.arr_char J.cod_char J.dom_char J.seq_char mkCone_def
\<chi>.is_natural_1 \<chi>.is_natural_2 \<chi>.A.map_simp map_def
by (metis (no_types, lifting) J.Zero_not_eq_j0 J.dom_simp(2))
ultimately have "J.arr j \<Longrightarrow> mkCone (\<chi> J.Zero) j = \<chi> j"
using J.arr_char by auto
thus "mkCone (\<chi> J.Zero) j = \<chi> j"
using mkCone_def \<chi>.is_extensional by fastforce
qed
qed
end
locale equalizer_cone =
J: parallel_pair +
C: category C +
D: parallel_pair_diagram C f0 f1 +
limit_cone J.comp C D.map "C.dom e" "D.mkCone e"
for C :: "'c comp" (infixr "\<cdot>" 55)
and f0 :: 'c
and f1 :: 'c
and e :: 'c
begin
lemma equalizes:
shows "D.is_equalized_by e"
proof
show "C.seq f0 e"
proof (intro C.seqI)
show "C.arr e" using ide_apex C.arr_dom_iff_arr by fastforce
show "C.arr f0"
using D.map_simp D.preserves_arr J.arr_char by metis
show "C.dom f0 = C.cod e"
using J.arr_char J.ide_char D.mkCone_def D.map_simp preserves_cod [of J.Zero]
by auto
qed
show "f0 \<cdot> e = f1 \<cdot> e"
using D.map_simp D.mkCone_def J.arr_char naturality [of J.j0] naturality [of J.j1]
by force
qed
lemma is_universal':
assumes "D.is_equalized_by e'"
shows "\<exists>!h. \<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright> \<and> e \<cdot> h = e'"
proof -
have "D.cone (C.dom e') (D.mkCone e')"
using assms D.cone_mkCone by blast
moreover have 0: "D.cone (C.dom e) (D.mkCone e)" ..
ultimately have 1: "\<exists>!h. \<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright> \<and>
D.cones_map h (D.mkCone e) = D.mkCone e'"
using is_universal [of "C.dom e'" "D.mkCone e'"] by auto
have 2: "\<And>h. \<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright> \<Longrightarrow>
D.cones_map h (D.mkCone e) = D.mkCone e' \<longleftrightarrow> e \<cdot> h = e'"
proof -
fix h
assume h: "\<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright>"
show "D.cones_map h (D.mkCone e) = D.mkCone e' \<longleftrightarrow> e \<cdot> h = e'"
proof
assume 3: "D.cones_map h (D.mkCone e) = D.mkCone e'"
show "e \<cdot> h = e'"
proof -
have "e' = D.mkCone e' J.Zero"
using D.mkCone_def J.arr_char by simp
also have "... = D.cones_map h (D.mkCone e) J.Zero"
using 3 by simp
also have "... = e \<cdot> h"
using 0 h D.mkCone_def J.arr_char by auto
finally show ?thesis by auto
qed
next
assume e': "e \<cdot> h = e'"
show "D.cones_map h (D.mkCone e) = D.mkCone e'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map h (D.mkCone e) j = D.mkCone e' j"
using h cone_axioms D.mkCone_def by auto
moreover have "j = J.Zero \<Longrightarrow> D.cones_map h (D.mkCone e) j = D.mkCone e' j"
using h e' cone_\<chi> D.mkCone_def J.arr_char [of J.Zero] by force
moreover have
"J.arr j \<and> j \<noteq> J.Zero \<Longrightarrow> D.cones_map h (D.mkCone e) j = D.mkCone e' j"
using C.comp_assoc D.mkCone_def cone_\<chi> e' h by auto
ultimately show "D.cones_map h (D.mkCone e) j = D.mkCone e' j" by blast
qed
qed
qed
thus ?thesis using 1 by blast
qed
lemma induced_arrowI':
assumes "D.is_equalized_by e'"
shows "\<guillemotleft>induced_arrow (C.dom e') (D.mkCone e') : C.dom e' \<rightarrow> C.dom e\<guillemotright>"
and "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') = e'"
proof -
interpret A': constant_functor J.comp C \<open>C.dom e'\<close>
using assms by (unfold_locales, auto)
have cone: "D.cone (C.dom e') (D.mkCone e')"
using assms D.cone_mkCone [of e'] by blast
have "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') =
D.cones_map (induced_arrow (C.dom e') (D.mkCone e')) (D.mkCone e) J.Zero"
using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\<chi> by force
also have "... = e'"
proof -
have "D.cones_map (induced_arrow (C.dom e') (D.mkCone e')) (D.mkCone e) =
D.mkCone e'"
using cone induced_arrowI by blast
thus ?thesis
using J.arr_char D.mkCone_def by simp
qed
finally have 1: "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') = e'"
by auto
show "\<guillemotleft>induced_arrow (C.dom e') (D.mkCone e') : C.dom e' \<rightarrow> C.dom e\<guillemotright>"
using 1 cone induced_arrowI by simp
show "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') = e'"
using 1 cone induced_arrowI by simp
qed
end
context category
begin
definition has_as_equalizer
where "has_as_equalizer f0 f1 e \<equiv> par f0 f1 \<and> parallel_pair_diagram.has_as_equalizer C f0 f1 e"
definition has_equalizers
where "has_equalizers = (\<forall>f0 f1. par f0 f1 \<longrightarrow> (\<exists>e. has_as_equalizer f0 f1 e))"
lemma has_as_equalizerI [intro]:
assumes "par f g" and "seq f e" and "f \<cdot> e = g \<cdot> e"
and "\<And>e'. \<lbrakk>seq f e'; f \<cdot> e' = g \<cdot> e'\<rbrakk> \<Longrightarrow> \<exists>!h. e \<cdot> h = e'"
shows "has_as_equalizer f g e"
proof (unfold has_as_equalizer_def, intro conjI)
show "arr f" and "arr g" and "dom f = dom g" and "cod f = cod g"
using assms(1) by auto
interpret J: parallel_pair .
interpret D: parallel_pair_diagram C f g
using assms(1) by unfold_locales
show "D.has_as_equalizer e"
proof -
let ?\<chi> = "D.mkCone e"
let ?a = "dom e"
interpret \<chi>: cone J.comp C D.map ?a ?\<chi>
using assms(2-3) D.cone_mkCone [of e] by simp
interpret \<chi>: limit_cone J.comp C D.map ?a ?\<chi>
proof
fix a' \<chi>'
assume \<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J.comp C D.map a' \<chi>'
using \<chi>' by blast
have "seq f (\<chi>' J.Zero)"
using J.ide_char J.arr_char \<chi>'.preserves_hom
by (metis (no_types, lifting) D.map_simp(3) \<chi>'.is_natural_1
\<chi>'.natural_transformation_axioms natural_transformation.preserves_reflects_arr
parallel_pair.dom_simp(3))
moreover have "f \<cdot> (\<chi>' J.Zero) = g \<cdot> (\<chi>' J.Zero)"
using \<chi>' D.is_equalized_by_cone by blast
ultimately have 1: "\<exists>!h. e \<cdot> h = \<chi>' J.Zero"
using assms by blast
obtain h where h: "e \<cdot> h = \<chi>' J.Zero"
using 1 by blast
have 2: "D.is_equalized_by e"
using assms(2-3) by blast
have "\<guillemotleft>h : a' \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h (D.mkCone e) = \<chi>'"
proof
show 3: "\<guillemotleft>h : a' \<rightarrow> dom e\<guillemotright>"
using h \<chi>'.preserves_dom
by (metis (no_types, lifting) \<chi>'.component_in_hom \<open>seq f (\<chi>' J.Zero)\<close>
category.has_codomain_iff_arr category.seqE category_axioms cod_in_codomains
domains_comp in_hom_def parallel_pair.arr_char)
show "D.cones_map h (D.mkCone e) = \<chi>'"
proof
fix j
have "D.cone (cod h) (D.mkCone e)"
using 2 3 D.cone_mkCone by auto
thus "D.cones_map h (D.mkCone e) j = \<chi>' j"
using h 2 3 D.cone_mkCone [of e] J.arr_char D.mkCone_def comp_assoc
apply (cases "J.arr j")
apply simp_all
apply (metis (no_types, lifting) D.mkCone_cone \<chi>')
using \<chi>'.is_extensional
by presburger
qed
qed
hence "\<exists>h. \<guillemotleft>h : a' \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h (D.mkCone e) = \<chi>'"
by blast
moreover have "\<And>h'. \<guillemotleft>h' : a' \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h' (D.mkCone e) = \<chi>' \<Longrightarrow> h' = h"
proof (elim conjE)
fix h'
assume h': "\<guillemotleft>h' : a' \<rightarrow> dom e\<guillemotright>"
assume eq: "D.cones_map h' (D.mkCone e) = \<chi>'"
have "e \<cdot> h' = \<chi>' J.Zero"
using eq D.cone_mkCone D.mkCone_def \<chi>'.preserves_reflects_arr \<chi>.cone_axioms
\<open>seq f (\<chi>' J.Zero)\<close> eq h' in_homE mem_Collect_eq restrict_apply seqE
apply simp
by fastforce
moreover have "\<exists>!h. e \<cdot> h = \<chi>' J.Zero"
using assms(2,4) 1 category.seqI by blast
ultimately show "h' = h"
using h by auto
qed
ultimately show "\<exists>!h. \<guillemotleft>h : a' \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h (D.mkCone e) = \<chi>'"
by blast
qed
show "D.has_as_equalizer e"
using assms \<chi>.limit_cone_axioms by blast
qed
qed
end
section "Limits by Products and Equalizers"
text\<open>
A category with equalizers has limits of shape @{term J} if it has products
indexed by the set of arrows of @{term J} and the set of objects of @{term J}.
The proof is patterned after \cite{MacLane}, Theorem 2, page 109:
\begin{quotation}
\noindent
``The limit of \<open>F: J \<rightarrow> C\<close> is the equalizer \<open>e\<close>
of \<open>f, g: \<Pi>\<^sub>i F\<^sub>i \<rightarrow> \<Pi>\<^sub>u F\<^sub>c\<^sub>o\<^sub>d \<^sub>u (u \<in> arr J, i \<in> J)\<close>
where \<open>p\<^sub>u f = p\<^sub>c\<^sub>o\<^sub>d \<^sub>u\<close>, \<open>p\<^sub>u g = F\<^sub>u o p\<^sub>d\<^sub>o\<^sub>m \<^sub>u\<close>;
the limiting cone \<open>\<mu>\<close> is \<open>\<mu>\<^sub>j = p\<^sub>j e\<close>, for \<open>j \<in> J\<close>.''
\end{quotation}
\<close>
locale category_with_equalizers =
category C
for C :: "'c comp" (infixr "\<cdot>" 55) +
assumes has_equalizers: "has_equalizers"
begin
lemma has_limits_if_has_products:
fixes J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
assumes "category J" and "has_products (Collect (partial_magma.ide J))"
and "has_products (Collect (partial_magma.arr J))"
shows "has_limits_of_shape J"
proof (unfold has_limits_of_shape_def)
interpret J: category J using assms(1) by auto
have "\<And>D. diagram J C D \<Longrightarrow> (\<exists>a \<chi>. limit_cone J C D a \<chi>)"
proof -
fix D
assume D: "diagram J C D"
interpret D: diagram J C D using D by auto
text\<open>
First, construct the two required products and their cones.
\<close>
interpret Obj: discrete_category \<open>Collect J.ide\<close> J.null
using J.not_arr_null J.ideD(1) mem_Collect_eq by (unfold_locales, blast)
interpret \<Delta>o: discrete_diagram_from_map \<open>Collect J.ide\<close> C D J.null
using D.preserves_ide by (unfold_locales, auto)
have "\<exists>p. has_as_product Obj.comp \<Delta>o.map p"
using assms(2) \<Delta>o.diagram_axioms has_products_def Obj.arr_char
by (metis (no_types, lifting) Collect_cong \<Delta>o.discrete_diagram_axioms mem_Collect_eq)
from this obtain \<Pi>o \<pi>o where \<pi>o: "product_cone Obj.comp C \<Delta>o.map \<Pi>o \<pi>o"
using ex_productE [of Obj.comp \<Delta>o.map] by auto
interpret \<pi>o: product_cone Obj.comp C \<Delta>o.map \<Pi>o \<pi>o using \<pi>o by auto
have \<pi>o_in_hom: "\<And>j. Obj.arr j \<Longrightarrow> \<guillemotleft>\<pi>o j : \<Pi>o \<rightarrow> D j\<guillemotright>"
using \<pi>o.preserves_dom \<pi>o.preserves_cod \<Delta>o.map_def by auto
interpret Arr: discrete_category \<open>Collect J.arr\<close> J.null
using J.not_arr_null by (unfold_locales, blast)
interpret \<Delta>a: discrete_diagram_from_map \<open>Collect J.arr\<close> C \<open>D o J.cod\<close> J.null
by (unfold_locales, auto)
have "\<exists>p. has_as_product Arr.comp \<Delta>a.map p"
using assms(3) has_products_def [of "Collect J.arr"] \<Delta>a.discrete_diagram_axioms
by blast
from this obtain \<Pi>a \<pi>a where \<pi>a: "product_cone Arr.comp C \<Delta>a.map \<Pi>a \<pi>a"
using ex_productE [of Arr.comp \<Delta>a.map] by auto
interpret \<pi>a: product_cone Arr.comp C \<Delta>a.map \<Pi>a \<pi>a using \<pi>a by auto
have \<pi>a_in_hom: "\<And>j. Arr.arr j \<Longrightarrow> \<guillemotleft>\<pi>a j : \<Pi>a \<rightarrow> D (J.cod j)\<guillemotright>"
using \<pi>a.preserves_cod \<pi>a.preserves_dom \<Delta>a.map_def by auto
text\<open>
Next, construct a parallel pair of arrows \<open>f, g: \<Pi>o \<rightarrow> \<Pi>a\<close>
that expresses the commutativity constraints imposed by the diagram.
\<close>
interpret \<Pi>o: constant_functor Arr.comp C \<Pi>o
using \<pi>o.ide_apex by (unfold_locales, auto)
let ?\<chi> = "\<lambda>j. if Arr.arr j then \<pi>o (J.cod j) else null"
interpret \<chi>: cone Arr.comp C \<Delta>a.map \<Pi>o ?\<chi>
using \<pi>o.ide_apex \<pi>o_in_hom \<Delta>a.map_def \<Delta>o.map_def \<Delta>o.is_discrete \<pi>o.is_natural_2
comp_cod_arr
by (unfold_locales, auto)
let ?f = "\<pi>a.induced_arrow \<Pi>o ?\<chi>"
have f_in_hom: "\<guillemotleft>?f : \<Pi>o \<rightarrow> \<Pi>a\<guillemotright>"
using \<chi>.cone_axioms \<pi>a.induced_arrowI by blast
have f_map: "\<Delta>a.cones_map ?f \<pi>a = ?\<chi>"
using \<chi>.cone_axioms \<pi>a.induced_arrowI by blast
have ff: "\<And>j. J.arr j \<Longrightarrow> \<pi>a j \<cdot> ?f = \<pi>o (J.cod j)"
using \<chi>.component_in_hom \<pi>a.induced_arrowI' \<pi>o.ide_apex by auto
let ?\<chi>' = "\<lambda>j. if Arr.arr j then D j \<cdot> \<pi>o (J.dom j) else null"
interpret \<chi>': cone Arr.comp C \<Delta>a.map \<Pi>o ?\<chi>'
using \<pi>o.ide_apex \<pi>o_in_hom \<Delta>o.map_def \<Delta>a.map_def comp_arr_dom comp_cod_arr
by (unfold_locales, auto)
let ?g = "\<pi>a.induced_arrow \<Pi>o ?\<chi>'"
have g_in_hom: "\<guillemotleft>?g : \<Pi>o \<rightarrow> \<Pi>a\<guillemotright>"
using \<chi>'.cone_axioms \<pi>a.induced_arrowI by blast
have g_map: "\<Delta>a.cones_map ?g \<pi>a = ?\<chi>'"
using \<chi>'.cone_axioms \<pi>a.induced_arrowI by blast
have gg: "\<And>j. J.arr j \<Longrightarrow> \<pi>a j \<cdot> ?g = D j \<cdot> \<pi>o (J.dom j)"
using \<chi>'.component_in_hom \<pi>a.induced_arrowI' \<pi>o.ide_apex by force
interpret PP: parallel_pair_diagram C ?f ?g
using f_in_hom g_in_hom
by (elim in_homE, unfold_locales, auto)
from PP.is_parallel obtain e where equ: "PP.has_as_equalizer e"
using has_equalizers has_equalizers_def has_as_equalizer_def by blast
interpret EQU: limit_cone PP.J.comp C PP.map \<open>dom e\<close> \<open>PP.mkCone e\<close>
using equ by auto
interpret EQU: equalizer_cone C ?f ?g e ..
text\<open>
An arrow @{term h} with @{term "cod h = \<Pi>o"} equalizes @{term f} and @{term g}
if and only if it satisfies the commutativity condition required for a cone over
@{term D}.
\<close>
have E: "\<And>h. \<guillemotleft>h : dom h \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow>
?f \<cdot> h = ?g \<cdot> h \<longleftrightarrow> (\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h)"
proof
fix h
assume h: "\<guillemotleft>h : dom h \<rightarrow> \<Pi>o\<guillemotright>"
show "?f \<cdot> h = ?g \<cdot> h \<Longrightarrow> \<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h"
proof -
assume E: "?f \<cdot> h = ?g \<cdot> h"
have "\<And>j. J.arr j \<Longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h"
proof -
fix j
assume j: "J.arr j"
have "?\<chi> j \<cdot> h = \<Delta>a.cones_map ?f \<pi>a j \<cdot> h"
using j f_map by fastforce
also have "... = \<pi>a j \<cdot> ?f \<cdot> h"
using j f_in_hom \<Delta>a.map_def \<pi>a.cone_\<chi> comp_assoc by auto
also have "... = \<pi>a j \<cdot> ?g \<cdot> h"
using j E by simp
also have "... = \<Delta>a.cones_map ?g \<pi>a j \<cdot> h"
using j g_in_hom \<Delta>a.map_def \<pi>a.cone_\<chi> comp_assoc by auto
also have "... = ?\<chi>' j \<cdot> h"
using j g_map by force
finally show "?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h" by auto
qed
thus "\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h" by blast
qed
show "\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h \<Longrightarrow> ?f \<cdot> h = ?g \<cdot> h"
proof -
assume 1: "\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h"
have 2: "\<And>j. j \<in> Collect J.arr \<Longrightarrow> \<pi>a j \<cdot> ?f \<cdot> h = \<pi>a j \<cdot> ?g \<cdot> h"
proof -
fix j
assume j: "j \<in> Collect J.arr"
have "\<pi>a j \<cdot> ?f \<cdot> h = (\<pi>a j \<cdot> ?f) \<cdot> h"
using comp_assoc by simp
also have "... = ?\<chi> j \<cdot> h"
using ff j by force
also have "... = ?\<chi>' j \<cdot> h"
using 1 j by auto
also have "... = (\<pi>a j \<cdot> ?g) \<cdot> h"
using gg j by force
also have "... = \<pi>a j \<cdot> ?g \<cdot> h"
using comp_assoc by simp
finally show "\<pi>a j \<cdot> ?f \<cdot> h = \<pi>a j \<cdot> ?g \<cdot> h"
by auto
qed
show "C ?f h = C ?g h"
proof -
have "\<And>j. Arr.arr j \<Longrightarrow> \<guillemotleft>\<pi>a j \<cdot> ?f \<cdot> h : dom h \<rightarrow> \<Delta>a.map j\<guillemotright>"
using f_in_hom h \<pi>a_in_hom by (elim in_homE, auto)
hence 3: "\<exists>!k. \<guillemotleft>k : dom h \<rightarrow> \<Pi>a\<guillemotright> \<and> (\<forall>j. Arr.arr j \<longrightarrow> \<pi>a j \<cdot> k = \<pi>a j \<cdot> ?f \<cdot> h)"
using h \<pi>a \<pi>a.is_universal' [of "dom h" "\<lambda>j. \<pi>a j \<cdot> ?f \<cdot> h"] \<Delta>a.map_def
ide_dom [of h]
by blast
have 4: "\<And>P x x'. \<exists>!k. P k x \<Longrightarrow> P x x \<Longrightarrow> P x' x \<Longrightarrow> x' = x" by auto
let ?P = "\<lambda> k x. \<guillemotleft>k : dom h \<rightarrow> \<Pi>a\<guillemotright> \<and>
(\<forall>j. j \<in> Collect J.arr \<longrightarrow> \<pi>a j \<cdot> k = \<pi>a j \<cdot> x)"
have "?P (?g \<cdot> h) (?g \<cdot> h)"
using g_in_hom h by force
moreover have "?P (?f \<cdot> h) (?g \<cdot> h)"
using 2 f_in_hom g_in_hom h by force
ultimately show ?thesis
using 3 4 [of ?P "?f \<cdot> h" "?g \<cdot> h"] by auto
qed
qed
qed
have E': "\<And>e. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow>
?f \<cdot> e = ?g \<cdot> e \<longleftrightarrow>
(\<forall>j. J.arr j \<longrightarrow>
(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e)"
proof -
have 1: "\<And>e j. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow> J.arr j \<Longrightarrow>
?\<chi> j \<cdot> e = (D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e"
proof -
fix e j
assume e: "\<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright>"
assume j: "J.arr j"
have "\<guillemotleft>\<pi>o (J.cod j) \<cdot> e : dom e \<rightarrow> D (J.cod j)\<guillemotright>"
using e j \<pi>o_in_hom by auto
thus "?\<chi> j \<cdot> e = (D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e"
using j comp_arr_dom comp_cod_arr by (elim in_homE, auto)
qed
have 2: "\<And>e j. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow> J.arr j \<Longrightarrow> ?\<chi>' j \<cdot> e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e"
by (simp add: comp_assoc)
show "\<And>e. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow>
?f \<cdot> e = ?g \<cdot> e \<longleftrightarrow>
(\<forall>j. J.arr j \<longrightarrow>
(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e)"
using 1 2 E by presburger
qed
text\<open>
The composites of @{term e} with the projections from the product @{term \<Pi>o}
determine a limit cone @{term \<mu>} for @{term D}. The component of @{term \<mu>}
at an object @{term j} of @{term[source=true] J} is the composite @{term "C (\<pi>o j) e"}.
However, we need to extend @{term \<mu>} to all arrows @{term j} of @{term[source=true] J},
so the correct definition is @{term "\<mu> j = C (D j) (C (\<pi>o (J.dom j)) e)"}.
\<close>
have e_in_hom: "\<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright>"
using EQU.equalizes f_in_hom in_homI
by (metis (no_types, lifting) seqE in_homE)
have e_map: "C ?f e = C ?g e"
using EQU.equalizes f_in_hom in_homI by fastforce
interpret domE: constant_functor J C \<open>dom e\<close>
using e_in_hom by (unfold_locales, auto)
let ?\<mu> = "\<lambda>j. if J.arr j then D j \<cdot> \<pi>o (J.dom j) \<cdot> e else null"
have \<mu>: "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>?\<mu> j : dom e \<rightarrow> D (J.cod j)\<guillemotright>"
proof -
fix j
assume j: "J.arr j"
show "\<guillemotleft>?\<mu> j : dom e \<rightarrow> D (J.cod j)\<guillemotright>"
using j e_in_hom \<pi>o_in_hom [of "J.dom j"] by auto
qed
interpret \<mu>: cone J C D \<open>dom e\<close> ?\<mu>
using \<mu> comp_cod_arr e_in_hom e_map E'
apply unfold_locales
apply auto
by (metis D.as_nat_trans.is_natural_1 comp_assoc)
text\<open>
If @{term \<tau>} is any cone over @{term D} then @{term \<tau>} restricts to a cone over
@{term \<Delta>o} for which the induced arrow to @{term \<Pi>o} equalizes @{term f} and @{term g}.
\<close>
have R: "\<And>a \<tau>. cone J C D a \<tau> \<Longrightarrow>
cone Obj.comp C \<Delta>o.map a (\<Delta>o.mkCone \<tau>) \<and>
?f \<cdot> \<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)
= ?g \<cdot> \<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)"
proof -
fix a \<tau>
assume cone_\<tau>: "cone J C D a \<tau>"
interpret \<tau>: cone J C D a \<tau> using cone_\<tau> by auto
interpret A: constant_functor Obj.comp C a
using \<tau>.ide_apex by (unfold_locales, auto)
interpret \<tau>o: cone Obj.comp C \<Delta>o.map a \<open>\<Delta>o.mkCone \<tau>\<close>
using A.value_is_ide \<Delta>o.map_def comp_cod_arr comp_arr_dom
by (unfold_locales, auto)
let ?e = "\<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)"
have mkCone_\<tau>: "\<Delta>o.mkCone \<tau> \<in> \<Delta>o.cones a"
using \<tau>o.cone_axioms by auto
have e: "\<guillemotleft>?e : a \<rightarrow> \<Pi>o\<guillemotright>"
using mkCone_\<tau> \<pi>o.induced_arrowI by simp
have ee: "\<And>j. J.ide j \<Longrightarrow> \<pi>o j \<cdot> ?e = \<tau> j"
proof -
fix j
assume j: "J.ide j"
have "\<pi>o j \<cdot> ?e = \<Delta>o.cones_map ?e \<pi>o j"
using j e \<pi>o.cone_axioms by force
also have "... = \<Delta>o.mkCone \<tau> j"
using j mkCone_\<tau> \<pi>o.induced_arrowI [of "\<Delta>o.mkCone \<tau>" a] by fastforce
also have "... = \<tau> j"
using j by simp
finally show "\<pi>o j \<cdot> ?e = \<tau> j" by auto
qed
have "\<And>j. J.arr j \<Longrightarrow>
(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e) \<cdot> dom ?e = D j \<cdot> \<pi>o (J.dom j) \<cdot> ?e"
proof -
fix j
assume j: "J.arr j"
have 1: "\<guillemotleft>\<pi>o (J.cod j) : \<Pi>o \<rightarrow> D (J.cod j)\<guillemotright>" using j \<pi>o_in_hom by simp
have 2: "(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e) \<cdot> dom ?e
= D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e"
proof -
have "seq (D (J.cod j)) (\<pi>o (J.cod j))"
using j 1 by auto
moreover have "seq (\<pi>o (J.cod j)) ?e"
using j e by fastforce
ultimately show ?thesis using comp_arr_dom by auto
qed
also have 3: "... = \<pi>o (J.cod j) \<cdot> ?e"
using j e 1 comp_cod_arr by (elim in_homE, auto)
also have "... = D j \<cdot> \<pi>o (J.dom j) \<cdot> ?e"
using j e ee 2 3 \<tau>.naturality \<tau>.A.map_simp \<tau>.ide_apex comp_cod_arr by auto
finally show "(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e) \<cdot> dom ?e = D j \<cdot> \<pi>o (J.dom j) \<cdot> ?e"
by auto
qed
hence "C ?f ?e = C ?g ?e"
using E' \<pi>o.induced_arrowI \<tau>o.cone_axioms mem_Collect_eq by blast
thus "cone Obj.comp C \<Delta>o.map a (\<Delta>o.mkCone \<tau>) \<and> C ?f ?e = C ?g ?e"
using \<tau>o.cone_axioms by auto
qed
text\<open>
Finally, show that @{term \<mu>} is a limit cone.
\<close>
interpret \<mu>: limit_cone J C D \<open>dom e\<close> ?\<mu>
proof
fix a \<tau>
assume cone_\<tau>: "cone J C D a \<tau>"
interpret \<tau>: cone J C D a \<tau> using cone_\<tau> by auto
interpret A: constant_functor Obj.comp C a
using \<tau>.ide_apex by unfold_locales auto
have cone_\<tau>o: "cone Obj.comp C \<Delta>o.map a (\<Delta>o.mkCone \<tau>)"
using A.value_is_ide \<Delta>o.map_def D.preserves_ide comp_cod_arr comp_arr_dom
\<tau>.preserves_hom
by unfold_locales auto
show "\<exists>!h. \<guillemotleft>h : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h ?\<mu> = \<tau>"
proof
let ?e' = "\<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)"
have e'_in_hom: "\<guillemotleft>?e' : a \<rightarrow> \<Pi>o\<guillemotright>"
using cone_\<tau> R \<pi>o.induced_arrowI by auto
have e'_map: "?f \<cdot> ?e' = ?g \<cdot> ?e' \<and> \<Delta>o.cones_map ?e' \<pi>o = \<Delta>o.mkCone \<tau>"
using cone_\<tau> R \<pi>o.induced_arrowI [of "\<Delta>o.mkCone \<tau>" a] by auto
have equ: "PP.is_equalized_by ?e'"
using e'_map e'_in_hom f_in_hom seqI' by blast
let ?h = "EQU.induced_arrow a (PP.mkCone ?e')"
have h_in_hom: "\<guillemotleft>?h : a \<rightarrow> dom e\<guillemotright>"
using EQU.induced_arrowI PP.cone_mkCone [of ?e'] e'_in_hom equ by fastforce
have h_map: "PP.cones_map ?h (PP.mkCone e) = PP.mkCone ?e'"
using EQU.induced_arrowI [of "PP.mkCone ?e'" a] PP.cone_mkCone [of ?e']
e'_in_hom equ
by fastforce
have 3: "D.cones_map ?h ?\<mu> = \<tau>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map ?h ?\<mu> j = \<tau> j"
using h_in_hom \<mu>.cone_axioms cone_\<tau> \<tau>.is_extensional by force
moreover have "J.arr j \<Longrightarrow> D.cones_map ?h ?\<mu> j = \<tau> j"
proof -
fix j
assume j: "J.arr j"
have 1: "\<guillemotleft>\<pi>o (J.dom j) \<cdot> e : dom e \<rightarrow> D (J.dom j)\<guillemotright>"
using j e_in_hom \<pi>o_in_hom [of "J.dom j"] by auto
have "D.cones_map ?h ?\<mu> j = ?\<mu> j \<cdot> ?h"
using h_in_hom j \<mu>.cone_axioms by auto
also have "... = D j \<cdot> (\<pi>o (J.dom j) \<cdot> e) \<cdot> ?h"
using j comp_assoc by simp
also have "... = D j \<cdot> \<tau> (J.dom j)"
proof -
have "(\<pi>o (J.dom j) \<cdot> e) \<cdot> ?h = \<tau> (J.dom j)"
proof -
have "(\<pi>o (J.dom j) \<cdot> e) \<cdot> ?h = \<pi>o (J.dom j) \<cdot> e \<cdot> ?h"
using j 1 e_in_hom h_in_hom \<pi>o arrI comp_assoc by auto
also have "... = \<pi>o (J.dom j) \<cdot> ?e'"
using equ e'_in_hom EQU.induced_arrowI' [of ?e'] by auto
also have "... = \<Delta>o.cones_map ?e' \<pi>o (J.dom j)"
using j e'_in_hom \<pi>o.cone_axioms by auto
also have "... = \<tau> (J.dom j)"
using j e'_map by simp
finally show ?thesis by auto
qed
thus ?thesis by simp
qed
also have "... = \<tau> j"
using j \<tau>.is_natural_1 by simp
finally show "D.cones_map ?h ?\<mu> j = \<tau> j" by auto
qed
ultimately show "D.cones_map ?h ?\<mu> j = \<tau> j" by auto
qed
show "\<guillemotleft>?h : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map ?h ?\<mu> = \<tau>"
using h_in_hom 3 by simp
show "\<And>h'. \<guillemotleft>h' : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h' ?\<mu> = \<tau> \<Longrightarrow> h' = ?h"
proof -
fix h'
assume h': "\<guillemotleft>h' : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h' ?\<mu> = \<tau>"
have h'_in_hom: "\<guillemotleft>h' : a \<rightarrow> dom e\<guillemotright>" using h' by simp
have h'_map: "D.cones_map h' ?\<mu> = \<tau>" using h' by simp
show "h' = ?h"
proof -
have 1: "\<guillemotleft>e \<cdot> h' : a \<rightarrow> \<Pi>o\<guillemotright> \<and> ?f \<cdot> e \<cdot> h' = ?g \<cdot> e \<cdot> h' \<and>
\<Delta>o.cones_map (C e h') \<pi>o = \<Delta>o.mkCone \<tau>"
proof -
have 2: "\<guillemotleft>e \<cdot> h' : a \<rightarrow> \<Pi>o\<guillemotright>" using h'_in_hom e_in_hom by auto
moreover have "?f \<cdot> e \<cdot> h' = ?g \<cdot> e \<cdot> h'"
by (metis (no_types, lifting) EQU.equalizes comp_assoc)
moreover have "\<Delta>o.cones_map (e \<cdot> h') \<pi>o = \<Delta>o.mkCone \<tau>"
proof
have "\<Delta>o.cones_map (e \<cdot> h') \<pi>o = \<Delta>o.cones_map h' (\<Delta>o.cones_map e \<pi>o)"
using \<pi>o.cone_axioms e_in_hom h'_in_hom \<Delta>o.cones_map_comp [of e h']
by fastforce
fix j
have "\<not>Obj.arr j \<Longrightarrow> \<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j"
using 2 e_in_hom h'_in_hom \<pi>o.cone_axioms by auto
moreover have "Obj.arr j \<Longrightarrow> \<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j"
proof -
assume j: "Obj.arr j"
have "\<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<pi>o j \<cdot> e \<cdot> h'"
using 2 j \<pi>o.cone_axioms by auto
also have "... = (\<pi>o j \<cdot> e) \<cdot> h'"
using comp_assoc by auto
also have "... = \<Delta>o.mkCone ?\<mu> j \<cdot> h'"
using j e_in_hom \<pi>o_in_hom comp_ide_arr [of "D j" "\<pi>o j \<cdot> e"]
by fastforce
also have "... = \<Delta>o.mkCone \<tau> j"
using j h' \<mu>.cone_axioms mem_Collect_eq by auto
finally show "\<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j" by auto
qed
ultimately show "\<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j" by auto
qed
ultimately show ?thesis by auto
qed
have "\<guillemotleft>e \<cdot> h' : a \<rightarrow> \<Pi>o\<guillemotright>" using 1 by simp
moreover have "e \<cdot> h' = ?e'"
using 1 cone_\<tau>o e'_in_hom e'_map \<pi>o.is_universal \<pi>o by blast
ultimately show "h' = ?h"
using 1 h'_in_hom h'_map EQU.is_universal' [of "e \<cdot> h'"]
EQU.induced_arrowI' [of ?e'] equ
by (elim in_homE) auto
qed
qed
qed
qed
have "limit_cone J C D (dom e) ?\<mu>" ..
thus "\<exists>a \<mu>. limit_cone J C D a \<mu>" by auto
qed
thus "\<forall>D. diagram J C D \<longrightarrow> (\<exists>a \<mu>. limit_cone J C D a \<mu>)" by blast
qed
end
section "Limits in a Set Category"
text\<open>
In this section, we consider the special case of limits in a set category.
\<close>
locale diagram_in_set_category =
J: category J +
S: set_category S is_set +
diagram J S D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and S :: "'s comp" (infixr "\<cdot>" 55)
and is_set :: "'s set \<Rightarrow> bool"
and D :: "'j \<Rightarrow> 's"
begin
notation S.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text\<open>
An object @{term a} of a set category @{term[source=true] S} is a limit of a diagram in
@{term[source=true] S} if and only if there is a bijection between the set
@{term "S.hom S.unity a"} of points of @{term a} and the set of cones over the diagram
that have apex @{term S.unity}.
\<close>
lemma limits_are_sets_of_cones:
shows "has_as_limit a \<longleftrightarrow> S.ide a \<and> (\<exists>\<phi>. bij_betw \<phi> (S.hom S.unity a) (cones S.unity))"
proof
text\<open>
If \<open>has_limit a\<close>, then by the universal property of the limit cone,
composition in @{term[source=true] S} yields a bijection between @{term "S.hom S.unity a"}
and @{term "cones S.unity"}.
\<close>
assume a: "has_as_limit a"
hence "S.ide a"
using limit_cone_def cone.ide_apex by metis
from a obtain \<chi> where \<chi>: "limit_cone a \<chi>" by auto
interpret \<chi>: limit_cone J S D a \<chi> using \<chi> by auto
have "bij_betw (\<lambda>f. cones_map f \<chi>) (S.hom S.unity a) (cones S.unity)"
using \<chi>.bij_betw_hom_and_cones S.ide_unity by simp
thus "S.ide a \<and> (\<exists>\<phi>. bij_betw \<phi> (S.hom S.unity a) (cones S.unity))"
using \<open>S.ide a\<close> by blast
next
text\<open>
Conversely, an arbitrary bijection @{term \<phi>} between @{term "S.hom S.unity a"}
and cones unity extends pointwise to a natural bijection @{term "\<Phi> a'"} between
@{term "S.hom a' a"} and @{term "cones a'"}, showing that @{term a} is a limit.
In more detail, the hypotheses give us a correspondence between points of @{term a}
and cones with apex @{term "S.unity"}. We extend this to a correspondence between
functions to @{term a} and general cones, with each arrow from @{term a'} to @{term a}
determining a cone with apex @{term a'}. If @{term "f \<in> hom a' a"} then composition
with @{term f} takes each point @{term y} of @{term a'} to the point @{term "S f y"}
of @{term a}. To this we may apply the given bijection @{term \<phi>} to obtain
@{term "\<phi> (S f y) \<in> cones S.unity"}. The component @{term "\<phi> (S f y) j"} at @{term j}
of this cone is a point of @{term "S.cod (D j)"}. Thus, @{term "f \<in> hom a' a"} determines
a cone @{term \<chi>f} with apex @{term a'} whose component at @{term j} is the
unique arrow @{term "\<chi>f j"} of @{term[source=true] S} such that
@{term "\<chi>f j \<in> hom a' (cod (D j))"} and @{term "S (\<chi>f j) y = \<phi> (S f y) j"}
for all points @{term y} of @{term a'}.
The cone @{term \<chi>a} corresponding to @{term "a \<in> S.hom a a"} is then a limit cone.
\<close>
assume a: "S.ide a \<and> (\<exists>\<phi>. bij_betw \<phi> (S.hom S.unity a) (cones S.unity))"
hence ide_a: "S.ide a" by auto
show "has_as_limit a"
proof -
from a obtain \<phi> where \<phi>: "bij_betw \<phi> (S.hom S.unity a) (cones S.unity)" by blast
have X: "\<And>f j y. \<lbrakk> \<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>; J.arr j; \<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright> \<rbrakk>
\<Longrightarrow> \<guillemotleft>\<phi> (f \<cdot> y) j : S.unity \<rightarrow> S.cod (D j)\<guillemotright>"
proof -
fix f j y
assume f: "\<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>" and j: "J.arr j" and y: "\<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright>"
interpret \<chi>: cone J S D S.unity \<open>\<phi> (S f y)\<close>
using f y \<phi> bij_betw_imp_funcset funcset_mem by blast
show "\<guillemotleft>\<phi> (f \<cdot> y) j : S.unity \<rightarrow> S.cod (D j)\<guillemotright>" using j by auto
qed
text\<open>
We want to define the component @{term "\<chi>j \<in> S.hom (S.dom f) (S.cod (D j))"}
at @{term j} of a cone by specifying how it acts by composition on points
@{term "y \<in> S.hom S.unity (S.dom f)"}. We can do this because @{term[source=true] S}
is a set category.
\<close>
let ?P = "\<lambda>f j \<chi>j. \<guillemotleft>\<chi>j : S.dom f \<rightarrow> S.cod (D j)\<guillemotright> \<and>
(\<forall>y. \<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright> \<longrightarrow> \<chi>j \<cdot> y = \<phi> (f \<cdot> y) j)"
let ?\<chi> = "\<lambda>f j. if J.arr j then (THE \<chi>j. ?P f j \<chi>j) else S.null"
have \<chi>: "\<And>f j. \<lbrakk> \<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>; J.arr j \<rbrakk> \<Longrightarrow> ?P f j (?\<chi> f j)"
proof -
fix b f j
assume f: "\<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>" and j: "J.arr j"
interpret B: constant_functor J S \<open>S.dom f\<close>
using f by (unfold_locales) auto
have "(\<lambda>y. \<phi> (f \<cdot> y) j) \<in> S.hom S.unity (S.dom f) \<rightarrow> S.hom S.unity (S.cod (D j))"
using f j X Pi_I' by simp
hence "\<exists>!\<chi>j. ?P f j \<chi>j"
using f j S.fun_complete' by (elim S.in_homE) auto
thus "?P f j (?\<chi> f j)" using j theI' [of "?P f j"] by simp
qed
text\<open>
The arrows @{term "\<chi> f j"} are in fact the components of a cone with apex
@{term "S.dom f"}.
\<close>
have cone: "\<And>f. \<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright> \<Longrightarrow> cone (S.dom f) (?\<chi> f)"
proof -
fix f
assume f: "\<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>"
interpret B: constant_functor J S \<open>S.dom f\<close>
using f by unfold_locales auto
show "cone (S.dom f) (?\<chi> f)"
proof
show "\<And>j. \<not>J.arr j \<Longrightarrow> ?\<chi> f j = S.null" by simp
fix j
assume j: "J.arr j"
have 0: "\<guillemotleft>?\<chi> f j : S.dom f \<rightarrow> S.cod (D j)\<guillemotright>" using f j \<chi> by simp
show "S.dom (?\<chi> f j) = B.map (J.dom j)" using f j \<chi> by auto
show "S.cod (?\<chi> f j) = D (J.cod j)" using f j \<chi> by auto
have par2: "S.par (?\<chi> f (J.cod j) \<cdot> B.map j) (?\<chi> f j)"
using f j 0 \<chi> [of f "J.cod j"] by (elim S.in_homE, auto)
have nat: "\<And>y. \<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright> \<Longrightarrow>
(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = ?\<chi> f j \<cdot> y \<and>
(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f j \<cdot> y"
proof -
fix y
assume y: "\<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright>"
show "(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = ?\<chi> f j \<cdot> y \<and>
(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f j \<cdot> y"
proof
have 1: "\<phi> (f \<cdot> y) \<in> cones S.unity"
using f y \<phi> bij_betw_imp_funcset PiE
S.seqI S.cod_comp S.dom_comp mem_Collect_eq
by fastforce
interpret \<chi>: cone J S D S.unity \<open>\<phi> (f \<cdot> y)\<close>
using 1 by simp
show "(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = ?\<chi> f j \<cdot> y"
using J.arr_dom S.comp_assoc \<chi> \<chi>.is_natural_1 f j y by presburger
have "(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f (J.cod j) \<cdot> y"
using j B.map_simp par2 B.value_is_ide S.comp_arr_ide
by (metis (no_types, lifting))
also have "... = \<phi> (f \<cdot> y) (J.cod j)"
using f y \<chi> \<chi>.is_extensional by simp
also have "... = \<phi> (f \<cdot> y) j"
using j \<chi>.is_natural_2
by (metis J.arr_cod \<chi>.A.map_simp J.cod_cod)
also have "... = ?\<chi> f j \<cdot> y"
using f y \<chi> \<chi>.is_extensional by simp
finally show "(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f j \<cdot> y" by auto
qed
qed
show "D j \<cdot> ?\<chi> f (J.dom j) = ?\<chi> f j"
proof -
have "S.par (D j \<cdot> ?\<chi> f (J.dom j)) (?\<chi> f j)"
using f j 0 \<chi> [of f "J.dom j"] by (elim S.in_homE, auto)
thus ?thesis
using nat 0
- apply (intro S.arr_eqI' [of "D j \<cdot> ?\<chi> f (J.dom j)" "?\<chi> f j"])
+ apply (intro S.arr_eqI'\<^sub>S\<^sub>C [of "D j \<cdot> ?\<chi> f (J.dom j)" "?\<chi> f j"])
apply force
by auto
qed
show "?\<chi> f (J.cod j) \<cdot> B.map j = ?\<chi> f j"
using par2 nat 0 f j \<chi>
- apply (intro S.arr_eqI' [of "?\<chi> f (J.cod j) \<cdot> B.map j" "?\<chi> f j"])
+ apply (intro S.arr_eqI'\<^sub>S\<^sub>C [of "?\<chi> f (J.cod j) \<cdot> B.map j" "?\<chi> f j"])
apply force
by (metis (no_types, lifting) S.in_homE)
qed
qed
interpret \<chi>a: cone J S D a \<open>?\<chi> a\<close> using a cone [of a] by fastforce
text\<open>
Finally, show that \<open>\<chi> a\<close> is a limit cone.
\<close>
interpret \<chi>a: limit_cone J S D a \<open>?\<chi> a\<close>
proof
fix a' \<chi>'
assume cone_\<chi>': "cone a' \<chi>'"
interpret \<chi>': cone J S D a' \<chi>' using cone_\<chi>' by auto
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> cones_map f (?\<chi> a) = \<chi>'"
proof
let ?\<psi> = "inv_into (S.hom S.unity a) \<phi>"
have \<psi>: "?\<psi> \<in> cones S.unity \<rightarrow> S.hom S.unity a"
using \<phi> bij_betw_inv_into bij_betwE by blast
let ?P = "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and>
(\<forall>y. y \<in> S.hom S.unity a' \<longrightarrow> f \<cdot> y = ?\<psi> (cones_map y \<chi>'))"
have 1: "\<exists>!f. ?P f"
proof -
have "(\<lambda>y. ?\<psi> (cones_map y \<chi>')) \<in> S.hom S.unity a' \<rightarrow> S.hom S.unity a"
proof
fix x
assume "x \<in> S.hom S.unity a'"
hence "\<guillemotleft>x : S.unity \<rightarrow> a'\<guillemotright>" by simp
hence "cones_map x \<in> cones a' \<rightarrow> cones S.unity"
using cones_map_mapsto [of x] by (elim S.in_homE) auto
hence "cones_map x \<chi>' \<in> cones S.unity"
using cone_\<chi>' by blast
thus "?\<psi> (cones_map x \<chi>') \<in> S.hom S.unity a"
using \<psi> by auto
qed
thus ?thesis
using S.fun_complete' a \<chi>'.ide_apex by simp
qed
let ?f = "THE f. ?P f"
have f: "?P ?f" using 1 theI' [of ?P] by simp
have f_in_hom: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright>" using f by simp
have f_map: "cones_map ?f (?\<chi> a) = \<chi>'"
proof -
have 1: "cone a' (cones_map ?f (?\<chi> a))"
proof -
have "cones_map ?f \<in> cones a \<rightarrow> cones a'"
using f_in_hom cones_map_mapsto [of ?f] by (elim S.in_homE) auto
hence "cones_map ?f (?\<chi> a) \<in> cones a'"
using \<chi>a.cone_axioms by blast
thus ?thesis by simp
qed
interpret f\<chi>a: cone J S D a' \<open>cones_map ?f (?\<chi> a)\<close>
using 1 by simp
show ?thesis
proof
fix j
have "\<not>J.arr j \<Longrightarrow> cones_map ?f (?\<chi> a) j = \<chi>' j"
using 1 \<chi>'.is_extensional f\<chi>a.is_extensional by presburger
moreover have "J.arr j \<Longrightarrow> cones_map ?f (?\<chi> a) j = \<chi>' j"
proof -
assume j: "J.arr j"
show "cones_map ?f (?\<chi> a) j = \<chi>' j"
- proof (intro S.arr_eqI' [of "cones_map ?f (?\<chi> a) j" "\<chi>' j"])
+ proof (intro S.arr_eqI'\<^sub>S\<^sub>C [of "cones_map ?f (?\<chi> a) j" "\<chi>' j"])
show par: "S.par (cones_map ?f (?\<chi> a) j) (\<chi>' j)"
using j \<chi>'.preserves_cod \<chi>'.preserves_dom \<chi>'.preserves_reflects_arr
f\<chi>a.preserves_cod f\<chi>a.preserves_dom f\<chi>a.preserves_reflects_arr
by presburger
fix y
assume "\<guillemotleft>y : S.unity \<rightarrow> S.dom (cones_map ?f (?\<chi> a) j)\<guillemotright>"
hence y: "\<guillemotleft>y : S.unity \<rightarrow> a'\<guillemotright>"
using j f\<chi>a.preserves_dom by simp
have 1: "\<guillemotleft>?\<chi> a j : a \<rightarrow> D (J.cod j)\<guillemotright>"
using j \<chi>a.preserves_hom by force
have 2: "\<guillemotleft>?f \<cdot> y : S.unity \<rightarrow> a\<guillemotright>"
using f_in_hom y by blast
have "cones_map ?f (?\<chi> a) j \<cdot> y = (?\<chi> a j \<cdot> ?f) \<cdot> y"
proof -
have "S.cod ?f = a" using f_in_hom by blast
thus ?thesis using j \<chi>a.cone_axioms by simp
qed
also have "... = ?\<chi> a j \<cdot> ?f \<cdot> y"
using 1 j y f_in_hom S.comp_assoc S.seqI' by blast
also have "... = \<phi> (a \<cdot> ?f \<cdot> y) j"
using 1 2 ide_a f j y \<chi> [of a] by (simp add: S.ide_in_hom)
also have "... = \<phi> (?f \<cdot> y) j"
using a 2 y S.comp_cod_arr by (elim S.in_homE, auto)
also have "... = \<phi> (?\<psi> (cones_map y \<chi>')) j"
using j y f by simp
also have "... = cones_map y \<chi>' j"
proof -
have "cones_map y \<chi>' \<in> cones S.unity"
using cone_\<chi>' y cones_map_mapsto by force
hence "\<phi> (?\<psi> (cones_map y \<chi>')) = cones_map y \<chi>'"
using \<phi> bij_betw_inv_into_right [of \<phi>] by simp
thus ?thesis by auto
qed
also have "... = \<chi>' j \<cdot> y"
using cone_\<chi>' j y by auto
finally show "cones_map ?f (?\<chi> a) j \<cdot> y = \<chi>' j \<cdot> y"
by auto
qed
qed
ultimately show "cones_map ?f (?\<chi> a) j = \<chi>' j" by blast
qed
qed
show "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright> \<and> cones_map ?f (?\<chi> a) = \<chi>'"
using f_in_hom f_map by simp
show "\<And>f'. \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> cones_map f' (?\<chi> a) = \<chi>' \<Longrightarrow> f' = ?f"
proof -
fix f'
assume f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> cones_map f' (?\<chi> a) = \<chi>'"
have f'_in_hom: "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>" using f' by simp
have f'_map: "cones_map f' (?\<chi> a) = \<chi>'" using f' by simp
show "f' = ?f"
- proof (intro S.arr_eqI' [of f' ?f])
+ proof (intro S.arr_eqI'\<^sub>S\<^sub>C [of f' ?f])
show "S.par f' ?f"
using f_in_hom f'_in_hom by (elim S.in_homE, auto)
show "\<And>y'. \<guillemotleft>y' : S.unity \<rightarrow> S.dom f'\<guillemotright> \<Longrightarrow> f' \<cdot> y' = ?f \<cdot> y'"
proof -
fix y'
assume y': "\<guillemotleft>y' : S.unity \<rightarrow> S.dom f'\<guillemotright>"
have 0: "\<phi> (f' \<cdot> y') = cones_map y' \<chi>'"
proof
fix j
have 1: "\<guillemotleft>f' \<cdot> y' : S.unity \<rightarrow> a\<guillemotright>" using f'_in_hom y' by auto
hence 2: "\<phi> (f' \<cdot> y') \<in> cones S.unity"
using \<phi> bij_betw_imp_funcset [of \<phi> "S.hom S.unity a" "cones S.unity"]
by auto
interpret \<chi>'': cone J S D S.unity \<open>\<phi> (f' \<cdot> y')\<close> using 2 by auto
have "\<not>J.arr j \<Longrightarrow> \<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j"
using f' y' cone_\<chi>' \<chi>''.is_extensional mem_Collect_eq restrict_apply
by (elim S.in_homE, auto)
moreover have "J.arr j \<Longrightarrow> \<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j"
proof -
assume j: "J.arr j"
have 3: "\<guillemotleft>?\<chi> a j : a \<rightarrow> D (J.cod j)\<guillemotright>"
using j \<chi>a.preserves_hom by force
have "\<phi> (f' \<cdot> y') j = \<phi> (a \<cdot> f' \<cdot> y') j"
using a f' y' j S.comp_cod_arr by (elim S.in_homE, auto)
also have "... = ?\<chi> a j \<cdot> f' \<cdot> y'"
using 1 3 \<chi> [of a] a f' y' j by fastforce
also have "... = (?\<chi> a j \<cdot> f') \<cdot> y'"
using S.comp_assoc by simp
also have "... = cones_map f' (?\<chi> a) j \<cdot> y'"
using f' y' j \<chi>a.cone_axioms by auto
also have "... = \<chi>' j \<cdot> y'"
using f' by blast
also have "... = cones_map y' \<chi>' j"
using y' j cone_\<chi>' f' mem_Collect_eq restrict_apply by force
finally show "\<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j" by auto
qed
ultimately show "\<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j" by auto
qed
hence "f' \<cdot> y' = ?\<psi> (cones_map y' \<chi>')"
using \<phi> f'_in_hom y' S.comp_in_homI
bij_betw_inv_into_left [of \<phi> "S.hom S.unity a" "cones S.unity" "f' \<cdot> y'"]
by (elim S.in_homE, auto)
moreover have "?f \<cdot> y' = ?\<psi> (cones_map y' \<chi>')"
using \<phi> 0 1 f f_in_hom f'_in_hom y' S.comp_in_homI
bij_betw_inv_into_left [of \<phi> "S.hom S.unity a" "cones S.unity" "?f \<cdot> y'"]
by (elim S.in_homE, auto)
ultimately show "f' \<cdot> y' = ?f \<cdot> y'" by auto
qed
qed
qed
qed
qed
have "limit_cone a (?\<chi> a)" ..
thus ?thesis by auto
qed
qed
end
locale diagram_in_replete_set_category =
J: category J +
S: replete_set_category S +
diagram J S D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and S :: "'s comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 's"
begin
sublocale diagram_in_set_category J S S.setp D
..
end
context set_category
begin
text\<open>
A set category has an equalizer for any parallel pair of arrows.
\<close>
- lemma has_equalizers:
+ lemma has_equalizers\<^sub>S\<^sub>C:
shows "has_equalizers"
proof (unfold has_equalizers_def)
have "\<And>f0 f1. par f0 f1 \<Longrightarrow> \<exists>e. has_as_equalizer f0 f1 e"
proof -
fix f0 f1
assume par: "par f0 f1"
interpret J: parallel_pair .
interpret PP: parallel_pair_diagram S f0 f1
using par by unfold_locales auto
interpret PP: diagram_in_set_category J.comp S setp PP.map ..
text\<open>
Let @{term a} be the object corresponding to the set of all images of equalizing points
of @{term "dom f0"}, and let @{term e} be the inclusion of @{term a} in @{term "dom f0"}.
\<close>
let ?a = "mkIde (img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e})"
have 0: "{e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e} \<subseteq> hom unity (dom f0)"
by auto
hence 1: "img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e} \<subseteq> Univ"
using img_point_in_Univ by auto
have 2: "setp (img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e})"
proof -
have "setp (img ` hom unity (dom f0))"
using ide_dom par setp_img_points by blast
moreover have "img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e} \<subseteq>
img ` hom unity (dom f0)"
by blast
ultimately show ?thesis
by (meson setp_respects_subset)
qed
have ide_a: "ide ?a" using 1 2 ide_mkIde by auto
have set_a: "set ?a = img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using 1 2 set_mkIde by simp
have incl_in_a: "incl_in ?a (dom f0)"
proof -
have "ide (dom f0)"
using PP.is_parallel by simp
moreover have "set ?a \<subseteq> set (dom f0)"
using img_point_elem_set set_a by fastforce
ultimately show ?thesis
using incl_in_def \<open>ide ?a\<close> by simp
qed
text\<open>
Then @{term "set a"} is in bijective correspondence with @{term "PP.cones unity"}.
\<close>
let ?\<phi> = "\<lambda>t. PP.mkCone (mkPoint (dom f0) t)"
let ?\<psi> = "\<lambda>\<chi>. img (\<chi> (J.Zero))"
have bij: "bij_betw ?\<phi> (set ?a) (PP.cones unity)"
proof (intro bij_betwI)
show "?\<phi> \<in> set ?a \<rightarrow> PP.cones unity"
proof
fix t
assume t: "t \<in> set ?a"
hence 1: "t \<in> img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using set_a by blast
then have 2: "mkPoint (dom f0) t \<in> hom unity (dom f0)"
using mkPoint_in_hom imageE mem_Collect_eq mkPoint_img(2) by auto
with 1 have 3: "mkPoint (dom f0) t \<in> {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using mkPoint_img(2) by auto
then have "PP.is_equalized_by (mkPoint (dom f0) t)"
using CollectD par by fastforce
thus "PP.mkCone (mkPoint (dom f0) t) \<in> PP.cones unity"
using 2 PP.cone_mkCone [of "mkPoint (dom f0) t"] by auto
qed
show "?\<psi> \<in> PP.cones unity \<rightarrow> set ?a"
proof
fix \<chi>
assume \<chi>: "\<chi> \<in> PP.cones unity"
interpret \<chi>: cone J.comp S PP.map unity \<chi> using \<chi> by auto
have "\<chi> (J.Zero) \<in> hom unity (dom f0) \<and> f0 \<cdot> \<chi> (J.Zero) = f1 \<cdot> \<chi> (J.Zero)"
using \<chi> PP.map_def PP.is_equalized_by_cone J.arr_char by auto
hence "img (\<chi> (J.Zero)) \<in> set ?a"
using set_a by simp
thus "?\<psi> \<chi> \<in> set ?a" by blast
qed
show "\<And>t. t \<in> set ?a \<Longrightarrow> ?\<psi> (?\<phi> t) = t"
using set_a J.arr_char PP.mkCone_def imageE mem_Collect_eq mkPoint_img(2)
by auto
show "\<And>\<chi>. \<chi> \<in> PP.cones unity \<Longrightarrow> ?\<phi> (?\<psi> \<chi>) = \<chi>"
proof -
fix \<chi>
assume \<chi>: "\<chi> \<in> PP.cones unity"
interpret \<chi>: cone J.comp S PP.map unity \<chi> using \<chi> by auto
have 1: "\<chi> (J.Zero) \<in> hom unity (dom f0) \<and> f0 \<cdot> \<chi> (J.Zero) = f1 \<cdot> \<chi> (J.Zero)"
using \<chi> PP.map_def PP.is_equalized_by_cone J.arr_char by auto
hence "img (\<chi> (J.Zero)) \<in> set ?a"
using set_a by simp
hence "img (\<chi> (J.Zero)) \<in> set (dom f0)"
using incl_in_a incl_in_def by auto
hence "mkPoint (dom f0) (img (\<chi> J.Zero)) = \<chi> J.Zero"
using 1 mkPoint_img(2) by blast
hence "?\<phi> (?\<psi> \<chi>) = PP.mkCone (\<chi> J.Zero)" by simp
also have "... = \<chi>"
using \<chi> PP.mkCone_cone by simp
finally show "?\<phi> (?\<psi> \<chi>) = \<chi>" by auto
qed
qed
text\<open>
It follows that @{term a} is a limit of \<open>PP\<close>, and that the limit cone gives an
equalizer of @{term f0} and @{term f1}.
\<close>
have "PP.has_as_limit ?a"
proof -
have "\<exists>\<mu>. bij_betw \<mu> (hom unity ?a) (set ?a)"
using bij_betw_points_and_set ide_a by auto
from this obtain \<mu> where \<mu>: "bij_betw \<mu> (hom unity ?a) (set ?a)" by blast
have "bij_betw (?\<phi> o \<mu>) (hom unity ?a) (PP.cones unity)"
using bij \<mu> bij_betw_comp_iff by blast
hence "\<exists>\<phi>. bij_betw \<phi> (hom unity ?a) (PP.cones unity)" by auto
thus ?thesis
using ide_a PP.limits_are_sets_of_cones by simp
qed
from this obtain \<epsilon> where \<epsilon>: "limit_cone J.comp S PP.map ?a \<epsilon>" by auto
have "PP.has_as_equalizer (\<epsilon> J.Zero)"
proof -
interpret \<epsilon>: limit_cone J.comp S PP.map ?a \<epsilon> using \<epsilon> by auto
have "PP.mkCone (\<epsilon> (J.Zero)) = \<epsilon>"
using \<epsilon> PP.mkCone_cone \<epsilon>.cone_axioms by simp
moreover have "dom (\<epsilon> (J.Zero)) = ?a"
using J.ide_char \<epsilon>.preserves_hom \<epsilon>.A.map_def by simp
ultimately show ?thesis
using \<epsilon> by simp
qed
thus "\<exists>e. has_as_equalizer f0 f1 e"
using par has_as_equalizer_def by auto
qed
thus "\<forall>f0 f1. par f0 f1 \<longrightarrow> (\<exists>e. has_as_equalizer f0 f1 e)" by auto
qed
end
sublocale set_category \<subseteq> category_with_equalizers S
- apply unfold_locales using has_equalizers by auto
+ apply unfold_locales using has_equalizers\<^sub>S\<^sub>C by auto
context set_category
begin
text\<open>
The aim of the next results is to characterize the conditions under which a set
category has products. In a traditional development of category theory,
one shows that the category \textbf{Set} of \emph{all} sets has all small
(\emph{i.e.}~set-indexed) products. In the present context we do not have a
category of \emph{all} sets, but rather only a category of all sets with
elements at a particular type. Clearly, we cannot expect such a category
to have products indexed by arbitrarily large sets. The existence of
@{term I}-indexed products in a set category @{term[source=true] S} implies that the universe
\<open>S.Univ\<close> of @{term[source=true] S} must be large enough to admit the formation of
@{term I}-tuples of its elements. Conversely, for a set category @{term[source=true] S}
the ability to form @{term I}-tuples in @{term Univ} implies that
@{term[source=true] S} has @{term I}-indexed products. Below we make this precise by
defining the notion of when a set category @{term[source=true] S}
``admits @{term I}-indexed tupling'' and we show that @{term[source=true] S}
has @{term I}-indexed products if and only if it admits @{term I}-indexed tupling.
The definition of ``@{term[source=true] S} admits @{term I}-indexed tupling'' says that
there is an injective map, from the space of extensional functions from
@{term I} to @{term Univ}, to @{term Univ}. However for a convenient
statement and proof of the desired result, the definition of extensional
function from theory @{theory "HOL-Library.FuncSet"} needs to be modified.
The theory @{theory "HOL-Library.FuncSet"} uses the definite, but arbitrarily chosen value
@{term undefined} as the value to be assumed by an extensional function outside
of its domain. In the context of the \<open>set_category\<close>, though, it is
more natural to use \<open>S.unity\<close>, which is guaranteed to be an element of the
universe of @{term[source=true] S}, for this purpose. Doing things that way makes it
simpler to establish a bijective correspondence between cones over @{term D} with apex
@{term unity} and the set of extensional functions @{term d} that map
each arrow @{term j} of @{term J} to an element @{term "d j"} of @{term "set (D j)"}.
Possibly it makes sense to go back and make this change in \<open>set_category\<close>,
but that would mean completely abandoning @{theory "HOL-Library.FuncSet"} and essentially
introducing a duplicate version for use with \<open>set_category\<close>.
As a compromise, what I have done here is to locally redefine the few notions from
@{theory "HOL-Library.FuncSet"} that I need in order to prove the next set of results.
The redefined notions are primed to avoid confusion with the original versions.
\<close>
definition extensional'
where "extensional' A \<equiv> {f. \<forall>x. x \<notin> A \<longrightarrow> f x = unity}"
abbreviation PiE'
where "PiE' A B \<equiv> Pi A B \<inter> extensional' A"
abbreviation restrict'
where "restrict' f A \<equiv> \<lambda>x. if x \<in> A then f x else unity"
lemma extensional'I [intro]:
assumes "\<And>x. x \<notin> A \<Longrightarrow> f x = unity"
shows "f \<in> extensional' A"
using assms extensional'_def by auto
lemma extensional'_arb:
assumes "f \<in> extensional' A" and "x \<notin> A"
shows "f x = unity"
using assms extensional'_def by fast
lemma extensional'_monotone:
assumes "A \<subseteq> B"
shows "extensional' A \<subseteq> extensional' B"
using assms extensional'_arb by fastforce
lemma PiE'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> PiE' A B \<subseteq> PiE' A C"
by auto
end
locale discrete_diagram_in_set_category =
S: set_category S \<SS> +
discrete_diagram J S D +
diagram_in_set_category J S \<SS> D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and S :: "'s comp" (infixr "\<cdot>" 55)
and \<SS> :: "'s set \<Rightarrow> bool"
and D :: "'j \<Rightarrow> 's"
begin
text\<open>
For @{term D} a discrete diagram in a set category, there is a bijective correspondence
between cones over @{term D} with apex unity and the set of extensional functions @{term d}
that map each arrow @{term j} of @{term[source=true] J} to an element of
@{term "S.set (D j)"}.
\<close>
abbreviation I
where "I \<equiv> Collect J.arr"
definition funToCone
where "funToCone F \<equiv> \<lambda>j. if J.arr j then S.mkPoint (D j) (F j) else S.null"
definition coneToFun
where "coneToFun \<chi> \<equiv> \<lambda>j. if J.arr j then S.img (\<chi> j) else S.unity"
lemma funToCone_mapsto:
shows "funToCone \<in> S.PiE' I (S.set o D) \<rightarrow> cones S.unity"
proof
fix F
assume F: "F \<in> S.PiE' I (S.set o D)"
interpret U: constant_functor J S S.unity
apply unfold_locales using S.ide_unity by auto
have "cone S.unity (funToCone F)"
proof
show "\<And>j. \<not>J.arr j \<Longrightarrow> funToCone F j = S.null"
using funToCone_def by simp
fix j
assume j: "J.arr j"
have "funToCone F j = S.mkPoint (D j) (F j)"
using j funToCone_def by simp
moreover have "... \<in> S.hom S.unity (D j)"
using F j is_discrete S.img_mkPoint(1) [of "D j"] by force
ultimately have 2: "funToCone F j \<in> S.hom S.unity (D j)" by auto
show 3: "S.dom (funToCone F j) = U.map (J.dom j)"
using 2 j U.map_simp by auto
show 4: "S.cod (funToCone F j) = D (J.cod j)"
using 2 j is_discrete by auto
show "D j \<cdot> funToCone F (J.dom j) = funToCone F j"
using 2 j is_discrete S.comp_cod_arr by auto
show "funToCone F (J.cod j) \<cdot> (U.map j) = funToCone F j"
using 3 j is_discrete U.map_simp S.arr_dom_iff_arr S.comp_arr_dom U.preserves_arr
by (metis J.ide_char)
qed
thus "funToCone F \<in> cones S.unity" by auto
qed
lemma coneToFun_mapsto:
shows "coneToFun \<in> cones S.unity \<rightarrow> S.PiE' I (S.set o D)"
proof
fix \<chi>
assume \<chi>: "\<chi> \<in> cones S.unity"
interpret \<chi>: cone J S D S.unity \<chi> using \<chi> by auto
show "coneToFun \<chi> \<in> S.PiE' I (S.set o D)"
proof
show "coneToFun \<chi> \<in> Pi I (S.set o D)"
using S.mkPoint_img(1) coneToFun_def is_discrete \<chi>.component_in_hom
by (simp add: S.img_point_elem_set restrict_apply')
show "coneToFun \<chi> \<in> S.extensional' I"
by (metis S.extensional'I coneToFun_def mem_Collect_eq)
qed
qed
lemma funToCone_coneToFun:
assumes "\<chi> \<in> cones S.unity"
shows "funToCone (coneToFun \<chi>) = \<chi>"
proof
interpret \<chi>: cone J S D S.unity \<chi> using assms by auto
fix j
have "\<not>J.arr j \<Longrightarrow> funToCone (coneToFun \<chi>) j = \<chi> j"
using funToCone_def \<chi>.is_extensional by simp
moreover have "J.arr j \<Longrightarrow> funToCone (coneToFun \<chi>) j = \<chi> j"
using funToCone_def coneToFun_def S.mkPoint_img(2) is_discrete \<chi>.component_in_hom
by auto
ultimately show "funToCone (coneToFun \<chi>) j = \<chi> j" by blast
qed
lemma coneToFun_funToCone:
assumes "F \<in> S.PiE' I (S.set o D)"
shows "coneToFun (funToCone F) = F"
proof
fix i
have "i \<notin> I \<Longrightarrow> coneToFun (funToCone F) i = F i"
using assms coneToFun_def S.extensional'_arb [of F I i] by auto
moreover have "i \<in> I \<Longrightarrow> coneToFun (funToCone F) i = F i"
proof -
assume i: "i \<in> I"
have "coneToFun (funToCone F) i = S.img (funToCone F i)"
using i coneToFun_def by simp
also have "... = S.img (S.mkPoint (D i) (F i))"
using i funToCone_def by auto
also have "... = F i"
using assms i is_discrete S.img_mkPoint(2) by force
finally show "coneToFun (funToCone F) i = F i" by auto
qed
ultimately show "coneToFun (funToCone F) i = F i" by auto
qed
lemma bij_coneToFun:
shows "bij_betw coneToFun (cones S.unity) (S.PiE' I (S.set o D))"
using coneToFun_mapsto funToCone_mapsto funToCone_coneToFun coneToFun_funToCone
bij_betwI
by blast
lemma bij_funToCone:
shows "bij_betw funToCone (S.PiE' I (S.set o D)) (cones S.unity)"
using coneToFun_mapsto funToCone_mapsto funToCone_coneToFun coneToFun_funToCone
bij_betwI
by blast
end
context set_category
begin
text\<open>
A set category admits @{term I}-indexed tupling if there is an injective map that takes
each extensional function from @{term I} to @{term Univ} to an element of @{term Univ}.
\<close>
definition admits_tupling
where "admits_tupling I \<equiv> \<exists>\<pi>. \<pi> \<in> PiE' I (\<lambda>_. Univ) \<rightarrow> Univ \<and> inj_on \<pi> (PiE' I (\<lambda>_. Univ))"
lemma admits_tupling_monotone:
assumes "admits_tupling I" and "I' \<subseteq> I"
shows "admits_tupling I'"
proof -
from assms(1) obtain \<pi>
where \<pi>: "\<pi> \<in> PiE' I (\<lambda>_. Univ) \<rightarrow> Univ \<and> inj_on \<pi> (PiE' I (\<lambda>_. Univ))"
using admits_tupling_def by metis
have "\<pi> \<in> PiE' I' (\<lambda>_. Univ) \<rightarrow> Univ"
proof
fix f
assume f: "f \<in> PiE' I' (\<lambda>_. Univ)"
have "f \<in> PiE' I (\<lambda>_. Univ)"
- using assms(2) f extensional'_def [of I'] terminal_unity extensional'_monotone by auto
+ using assms(2) f extensional'_def [of I'] terminal_unity\<^sub>S\<^sub>C extensional'_monotone by auto
thus "\<pi> f \<in> Univ" using \<pi> by auto
qed
moreover have "inj_on \<pi> (PiE' I' (\<lambda>_. Univ))"
proof -
have 1: "\<And>F A A'. inj_on F A \<and> A' \<subseteq> A \<Longrightarrow> inj_on F A'"
using subset_inj_on by blast
moreover have "PiE' I' (\<lambda>_. Univ) \<subseteq> PiE' I (\<lambda>_. Univ)"
- using assms(2) extensional'_def [of I'] terminal_unity by auto
+ using assms(2) extensional'_def [of I'] terminal_unity\<^sub>S\<^sub>C by auto
ultimately show ?thesis using \<pi> assms(2) by blast
qed
ultimately show ?thesis using admits_tupling_def by metis
qed
lemma admits_tupling_respects_bij:
assumes "admits_tupling I" and "bij_betw \<phi> I I'"
shows "admits_tupling I'"
proof -
obtain \<pi> where \<pi>: "\<pi> \<in> (I \<rightarrow> Univ) \<inter> extensional' I \<rightarrow> Univ \<and>
inj_on \<pi> ((I \<rightarrow> Univ) \<inter> extensional' I)"
using assms(1) admits_tupling_def by metis
have inv: "bij_betw (inv_into I \<phi>) I' I"
using assms(2) bij_betw_inv_into by blast
let ?C = "\<lambda>f x. if x \<in> I then f (\<phi> x) else unity"
let ?\<pi>' = "\<lambda>f. \<pi> (?C f)"
have 1: "\<And>f. f \<in> (I' \<rightarrow> Univ) \<inter> extensional' I' \<Longrightarrow> ?C f \<in> (I \<rightarrow> Univ) \<inter> extensional' I"
using assms bij_betw_apply by fastforce
have "?\<pi>' \<in> (I' \<rightarrow> Univ) \<inter> extensional' I' \<rightarrow> Univ \<and>
inj_on ?\<pi>' ((I' \<rightarrow> Univ) \<inter> extensional' I')"
proof
show "(\<lambda>f. \<pi> (?C f)) \<in> (I' \<rightarrow> Univ) \<inter> extensional' I' \<rightarrow> Univ"
using 1 \<pi> by blast
show "inj_on ?\<pi>' ((I' \<rightarrow> Univ) \<inter> extensional' I')"
proof
fix f g
assume f: "f \<in> (I' \<rightarrow> Univ) \<inter> extensional' I'"
assume g: "g \<in> (I' \<rightarrow> Univ) \<inter> extensional' I'"
assume eq: "?\<pi>' f = ?\<pi>' g"
have f': "?C f \<in> (I \<rightarrow> Univ) \<inter> extensional' I"
using f 1 by simp
have g': "?C g \<in> (I \<rightarrow> Univ) \<inter> extensional' I"
using g 1 by simp
have 2: "?C f = ?C g"
using f' g' \<pi> eq by (simp add: inj_on_def)
show "f = g"
proof
fix x
show "f x = g x"
proof (cases "x \<in> I'")
show "x \<in> I' \<Longrightarrow> ?thesis"
using f g
by (metis (no_types, opaque_lifting) "2" assms(2) bij_betw_apply
bij_betw_inv_into_right inv)
show "x \<notin> I' \<Longrightarrow> ?thesis"
using f g by (metis IntD2 extensional'_arb)
qed
qed
qed
qed
thus ?thesis
using admits_tupling_def by blast
qed
end
context replete_set_category
begin
lemma has_products_iff_admits_tupling:
fixes I :: "'i set"
shows "has_products I \<longleftrightarrow> I \<noteq> UNIV \<and> admits_tupling I"
proof
text\<open>
If @{term[source=true] S} has @{term I}-indexed products, then for every @{term I}-indexed
discrete diagram @{term D} in @{term[source=true] S} there is an object @{term \<Pi>D}
of @{term[source=true] S} whose points are in bijective correspondence with the set of
cones over @{term D} with apex @{term unity}. In particular this is true for
the diagram @{term D} that assigns to each element of @{term I} the
``universal object'' @{term "mkIde Univ"}.
\<close>
assume has_products: "has_products I"
have I: "I \<noteq> UNIV" using has_products has_products_def by auto
interpret J: discrete_category I \<open>SOME x. x \<notin> I\<close>
using I someI_ex [of "\<lambda>x. x \<notin> I"] by (unfold_locales, auto)
let ?D = "\<lambda>i. mkIde Univ"
interpret D: discrete_diagram_from_map I S ?D \<open>SOME j. j \<notin> I\<close>
using J.not_arr_null J.arr_char ide_mkIde
by (unfold_locales, auto)
interpret D: discrete_diagram_in_set_category J.comp S \<open>\<lambda>A. A \<subseteq> Univ\<close> D.map ..
have "discrete_diagram J.comp S D.map" ..
from this obtain \<Pi>D \<chi> where \<chi>: "product_cone J.comp S D.map \<Pi>D \<chi>"
using has_products has_products_def [of I] ex_productE [of "J.comp" D.map]
D.diagram_axioms
by blast
interpret \<chi>: product_cone J.comp S D.map \<Pi>D \<chi>
using \<chi> by auto
have "D.has_as_limit \<Pi>D"
using \<chi>.limit_cone_axioms by auto
hence \<Pi>D: "ide \<Pi>D \<and> (\<exists>\<phi>. bij_betw \<phi> (hom unity \<Pi>D) (D.cones unity))"
using D.limits_are_sets_of_cones by simp
from this obtain \<phi> where \<phi>: "bij_betw \<phi> (hom unity \<Pi>D) (D.cones unity)"
by blast
have \<phi>': "inv_into (hom unity \<Pi>D) \<phi> \<in> D.cones unity \<rightarrow> hom unity \<Pi>D \<and>
inj_on (inv_into (hom unity \<Pi>D) \<phi>) (D.cones unity)"
using \<phi> bij_betw_inv_into bij_betw_imp_inj_on bij_betw_imp_funcset by blast
let ?\<pi> = "img o (inv_into (hom unity \<Pi>D) \<phi>) o D.funToCone"
have 1: "D.funToCone \<in> PiE' I (set o D.map) \<rightarrow> D.cones unity"
using D.funToCone_mapsto extensional'_def [of I] by auto
have 2: "inv_into (hom unity \<Pi>D) \<phi> \<in> D.cones unity \<rightarrow> hom unity \<Pi>D"
using \<phi>' by auto
have 3: "img \<in> hom unity \<Pi>D \<rightarrow> Univ"
using img_point_in_Univ by blast
have 4: "inj_on D.funToCone (PiE' I (set o D.map))"
proof -
have "D.I = I" by auto
thus ?thesis
using D.bij_funToCone bij_betw_imp_inj_on by auto
qed
have 5: "inj_on (inv_into (hom unity \<Pi>D) \<phi>) (D.cones unity)"
using \<phi>' by auto
have 6: "inj_on img (hom unity \<Pi>D)"
using \<Pi>D bij_betw_points_and_set bij_betw_imp_inj_on [of img "hom unity \<Pi>D" "set \<Pi>D"]
by simp
have "?\<pi> \<in> PiE' I (set o D.map) \<rightarrow> Univ"
using 1 2 3 by force
moreover have "inj_on ?\<pi> (PiE' I (set o D.map))"
proof -
have 7: "\<And>A B C D F G H. F \<in> A \<rightarrow> B \<and> G \<in> B \<rightarrow> C \<and> H \<in> C \<rightarrow> D
\<and> inj_on F A \<and> inj_on G B \<and> inj_on H C
\<Longrightarrow> inj_on (H o G o F) A"
by (simp add: Pi_iff inj_on_def)
show ?thesis
using 1 2 3 4 5 6 7 [of D.funToCone "PiE' I (set o D.map)" "D.cones unity"
"inv_into (hom unity \<Pi>D) \<phi>" "hom unity \<Pi>D"
img Univ]
by fastforce
qed
moreover have "PiE' I (set o D.map) = PiE' I (\<lambda>x. Univ)"
proof -
have "\<And>i. i \<in> I \<Longrightarrow> (set o D.map) i = Univ"
using J.arr_char D.map_def by simp
thus ?thesis by blast
qed
ultimately have "?\<pi> \<in> (PiE' I (\<lambda>x. Univ)) \<rightarrow> Univ \<and> inj_on ?\<pi> (PiE' I (\<lambda>x. Univ))"
by auto
thus "I \<noteq> UNIV \<and> admits_tupling I"
using I admits_tupling_def by auto
next
assume ex_\<pi>: "I \<noteq> UNIV \<and> admits_tupling I"
show "has_products I"
proof (unfold has_products_def)
from ex_\<pi> obtain \<pi>
where \<pi>: "\<pi> \<in> (PiE' I (\<lambda>x. Univ)) \<rightarrow> Univ \<and> inj_on \<pi> (PiE' I (\<lambda>x. Univ))"
using admits_tupling_def by metis
text\<open>
Given an @{term I}-indexed discrete diagram @{term D}, obtain the object @{term \<Pi>D}
of @{term[source=true] S} corresponding to the set @{term "\<pi> ` PiE I D"} of all
@{term "\<pi> d"} where \<open>d \<in> d \<in> J \<rightarrow>\<^sub>E Univ\<close> and @{term "d i \<in> D i"}
for all @{term "i \<in> I"}.
The elements of @{term \<Pi>D} are in bijective correspondence with the set of cones
over @{term D}, hence @{term \<Pi>D} is a limit of @{term D}.
\<close>
have "\<And>J D. discrete_diagram J S D \<and> Collect (partial_magma.arr J) = I
\<Longrightarrow> \<exists>\<Pi>D. has_as_product J D \<Pi>D"
proof
fix J :: "'i comp" and D
assume D: "discrete_diagram J S D \<and> Collect (partial_magma.arr J) = I"
interpret J: category J
using D discrete_diagram.axioms(1) by blast
interpret D: discrete_diagram J S D
using D by simp
interpret D: discrete_diagram_in_set_category J S \<open>\<lambda>A. A \<subseteq> Univ\<close> D ..
let ?\<Pi>D = "mkIde (\<pi> ` PiE' I (set o D))"
have 0: "ide ?\<Pi>D"
proof -
have "set o D \<in> I \<rightarrow> Pow Univ"
using Pow_iff incl_in_def o_apply elem_set_implies_incl_in subsetI Pi_I'
setp_set_ide
by (metis (mono_tags, lifting))
hence "\<pi> ` PiE' I (set o D) \<subseteq> Univ"
using \<pi> by blast
thus ?thesis using \<pi> ide_mkIde by simp
qed
hence set_\<Pi>D: "\<pi> ` PiE' I (set o D) = set ?\<Pi>D"
using 0 ide_in_hom arr_mkIde set_mkIde by auto
text\<open>
The elements of @{term \<Pi>D} are all values of the form @{term "\<pi> d"},
where @{term d} satisfies @{term "d i \<in> set (D i)"} for all @{term "i \<in> I"}.
Such @{term d} correspond bijectively to cones.
Since @{term \<pi>} is injective, the values @{term "\<pi> d"} correspond bijectively to cones.
\<close>
let ?\<phi> = "mkPoint ?\<Pi>D o \<pi> o D.coneToFun"
let ?\<phi>' = "D.funToCone o inv_into (PiE' I (set o D)) \<pi> o img"
have 1: "\<pi> \<in> PiE' I (set o D) \<rightarrow> set ?\<Pi>D \<and> inj_on \<pi> (PiE' I (set o D))"
proof -
have "PiE' I (set o D) \<subseteq> PiE' I (\<lambda>x. Univ)"
using setp_set_ide elem_set_implies_incl_in elem_set_implies_set_eq_singleton
incl_in_def PiE'_mono comp_apply subsetI
by (metis (no_types, lifting))
thus ?thesis using \<pi> subset_inj_on set_\<Pi>D Pi_I' imageI by fastforce
qed
have 2: "inv_into (PiE' I (set o D)) \<pi> \<in> set ?\<Pi>D \<rightarrow> PiE' I (set o D)"
proof
fix y
assume y: "y \<in> set ?\<Pi>D"
have "y \<in> \<pi> ` (PiE' I (set o D))" using y set_\<Pi>D by auto
thus "inv_into (PiE' I (set o D)) \<pi> y \<in> PiE' I (set o D)"
using inv_into_into [of y \<pi> "PiE' I (set o D)"] by simp
qed
have 3: "\<And>x. x \<in> set ?\<Pi>D \<Longrightarrow> \<pi> (inv_into (PiE' I (set o D)) \<pi> x) = x"
using set_\<Pi>D by (simp add: f_inv_into_f)
have 4: "\<And>d. d \<in> PiE' I (set o D) \<Longrightarrow> inv_into (PiE' I (set o D)) \<pi> (\<pi> d) = d"
using 1 by auto
have 5: "D.I = I"
using D by auto
have "bij_betw ?\<phi> (D.cones unity) (hom unity ?\<Pi>D)"
proof (intro bij_betwI)
show "?\<phi> \<in> D.cones unity \<rightarrow> hom unity ?\<Pi>D"
proof
fix \<chi>
assume \<chi>: "\<chi> \<in> D.cones unity"
show "?\<phi> \<chi> \<in> hom unity ?\<Pi>D"
using \<chi> 0 1 5 D.coneToFun_mapsto mkPoint_in_hom [of ?\<Pi>D]
by (simp, blast)
qed
show "?\<phi>' \<in> hom unity ?\<Pi>D \<rightarrow> D.cones unity"
proof
fix x
assume x: "x \<in> hom unity ?\<Pi>D"
hence "img x \<in> set ?\<Pi>D"
using img_point_elem_set by blast
hence "inv_into (PiE' I (set o D)) \<pi> (img x) \<in> Pi I (set \<circ> D) \<inter> extensional' I"
using 2 by blast
thus "?\<phi>' x \<in> D.cones unity"
using 5 D.funToCone_mapsto by auto
qed
show "\<And>x. x \<in> hom unity ?\<Pi>D \<Longrightarrow> ?\<phi> (?\<phi>' x) = x"
proof -
fix x
assume x: "x \<in> hom unity ?\<Pi>D"
show "?\<phi> (?\<phi>' x) = x"
proof -
have "D.coneToFun (D.funToCone (inv_into (PiE' I (set o D)) \<pi> (img x)))
= inv_into (PiE' I (set o D)) \<pi> (img x)"
using x 1 5 img_point_elem_set set_\<Pi>D D.coneToFun_funToCone by force
hence "\<pi> (D.coneToFun (D.funToCone (inv_into (PiE' I (set o D)) \<pi> (img x))))
= img x"
using x 3 img_point_elem_set set_\<Pi>D by force
thus ?thesis using x 0 mkPoint_img by auto
qed
qed
show "\<And>\<chi>. \<chi> \<in> D.cones unity \<Longrightarrow> ?\<phi>' (?\<phi> \<chi>) = \<chi>"
proof -
fix \<chi>
assume \<chi>: "\<chi> \<in> D.cones unity"
show "?\<phi>' (?\<phi> \<chi>) = \<chi>"
proof -
have "img (mkPoint ?\<Pi>D (\<pi> (D.coneToFun \<chi>))) = \<pi> (D.coneToFun \<chi>)"
using \<chi> 0 1 5 D.coneToFun_mapsto img_mkPoint(2) by blast
hence "inv_into (PiE' I (set o D)) \<pi> (img (mkPoint ?\<Pi>D (\<pi> (D.coneToFun \<chi>))))
= D.coneToFun \<chi>"
using \<chi> D.coneToFun_mapsto 4 5
by (metis (no_types, lifting) PiE)
hence "D.funToCone (inv_into (PiE' I (set o D)) \<pi>
(img (mkPoint ?\<Pi>D (\<pi> (D.coneToFun \<chi>)))))
= \<chi>"
using \<chi> D.funToCone_coneToFun by auto
thus ?thesis by auto
qed
qed
qed
hence "bij_betw (inv_into (D.cones unity) ?\<phi>) (hom unity ?\<Pi>D) (D.cones unity)"
using bij_betw_inv_into by blast
hence "\<exists>\<phi>. bij_betw \<phi> (hom unity ?\<Pi>D) (D.cones unity)" by blast
hence "D.has_as_limit ?\<Pi>D"
using \<open>ide ?\<Pi>D\<close> D.limits_are_sets_of_cones by simp
from this obtain \<chi> where \<chi>: "limit_cone J S D ?\<Pi>D \<chi>" by blast
interpret \<chi>: limit_cone J S D ?\<Pi>D \<chi> using \<chi> by auto
interpret P: product_cone J S D ?\<Pi>D \<chi>
using \<chi> D.product_coneI by blast
have "product_cone J S D ?\<Pi>D \<chi>" ..
thus "has_as_product J D ?\<Pi>D"
using has_as_product_def by auto
qed
thus "I \<noteq> UNIV \<and>
(\<forall>J D. discrete_diagram J S D \<and> Collect (partial_magma.arr J) = I
\<longrightarrow> (\<exists>\<Pi>D. has_as_product J D \<Pi>D))"
using ex_\<pi> by blast
qed
qed
end
context replete_set_category
begin
text\<open>
Characterization of the completeness properties enjoyed by a set category:
A set category @{term[source=true] S} has all limits at a type @{typ 'j},
if and only if @{term[source=true] S} admits @{term I}-indexed tupling
for all @{typ 'j}-sets @{term I} such that @{term "I \<noteq> UNIV"}.
\<close>
theorem has_limits_iff_admits_tupling:
shows "has_limits (undefined :: 'j) \<longleftrightarrow> (\<forall>I :: 'j set. I \<noteq> UNIV \<longrightarrow> admits_tupling I)"
proof
assume has_limits: "has_limits (undefined :: 'j)"
show "\<forall>I :: 'j set. I \<noteq> UNIV \<longrightarrow> admits_tupling I"
using has_limits has_products_if_has_limits has_products_iff_admits_tupling by blast
next
assume admits_tupling: "\<forall>I :: 'j set. I \<noteq> UNIV \<longrightarrow> admits_tupling I"
show "has_limits (undefined :: 'j)"
using has_limits_def admits_tupling has_products_iff_admits_tupling
by (metis category.axioms(1) category.ideD(1) has_limits_if_has_products
iso_tuple_UNIV_I mem_Collect_eq partial_magma.not_arr_null)
qed
end
section "Limits in Functor Categories"
text\<open>
In this section, we consider the special case of limits in functor categories,
with the objective of showing that limits in a functor category \<open>[A, B]\<close>
are given pointwise, and that \<open>[A, B]\<close> has all limits that @{term B} has.
\<close>
locale parametrized_diagram =
J: category J +
A: category A +
B: category B +
JxA: product_category J A +
binary_functor J A B D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and A :: "'a comp" (infixr "\<cdot>\<^sub>A" 55)
and B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and D :: "'j * 'a \<Rightarrow> 'b"
begin
(* Notation for A.in_hom and B.in_hom is being inherited, but from where? *)
notation J.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J _\<guillemotright>")
notation JxA.comp (infixr "\<cdot>\<^sub>J\<^sub>x\<^sub>A" 55)
notation JxA.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J\<^sub>x\<^sub>A _\<guillemotright>")
text\<open>
A choice of limit cone for each diagram \<open>D (-, a)\<close>, where @{term a}
is an object of @{term[source=true] A}, extends to a functor \<open>L: A \<rightarrow> B\<close>,
where the action of @{term L} on arrows of @{term[source=true] A} is determined by
universality.
\<close>
abbreviation L
where "L \<equiv> \<lambda>l \<chi>. \<lambda>a. if A.arr a then
limit_cone.induced_arrow J B (\<lambda>j. D (j, A.cod a))
(l (A.cod a)) (\<chi> (A.cod a))
(l (A.dom a)) (vertical_composite.map J B
(\<chi> (A.dom a)) (\<lambda>j. D (j, a)))
else B.null"
abbreviation P
where "P \<equiv> \<lambda>l \<chi>. \<lambda>a f. \<guillemotleft>f : l (A.dom a) \<rightarrow>\<^sub>B l (A.cod a)\<guillemotright> \<and>
diagram.cones_map J B (\<lambda>j. D (j, A.cod a)) f (\<chi> (A.cod a)) =
vertical_composite.map J B (\<chi> (A.dom a)) (\<lambda>j. D (j, a))"
lemma L_arr:
assumes "\<forall>a. A.ide a \<longrightarrow> limit_cone J B (\<lambda>j. D (j, a)) (l a) (\<chi> a)"
shows "\<And>a. A.arr a \<Longrightarrow> (\<exists>!f. P l \<chi> a f) \<and> P l \<chi> a (L l \<chi> a)"
proof
fix a
assume a: "A.arr a"
interpret \<chi>_dom_a: limit_cone J B \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>l (A.dom a)\<close> \<open>\<chi> (A.dom a)\<close>
using a assms by auto
interpret \<chi>_cod_a: limit_cone J B \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>l (A.cod a)\<close> \<open>\<chi> (A.cod a)\<close>
using a assms by auto
interpret Da: natural_transformation J B \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>\<lambda>j. D (j, A.cod a)\<close>
\<open>\<lambda>j. D (j, a)\<close>
using a fixing_arr_gives_natural_transformation_2 by simp
interpret Dao\<chi>_dom_a: vertical_composite J B
\<chi>_dom_a.A.map \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>\<lambda>j. D (j, A.cod a)\<close>
\<open>\<chi> (A.dom a)\<close> \<open>\<lambda>j. D (j, a)\<close> ..
interpret Dao\<chi>_dom_a: cone J B \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>l (A.dom a)\<close> Dao\<chi>_dom_a.map ..
show "P l \<chi> a (L l \<chi> a)"
using a Dao\<chi>_dom_a.cone_axioms \<chi>_cod_a.induced_arrowI [of Dao\<chi>_dom_a.map "l (A.dom a)"]
by auto
show "\<exists>!f. P l \<chi> a f"
using \<chi>_cod_a.is_universal Dao\<chi>_dom_a.cone_axioms by blast
qed
lemma L_ide:
assumes "\<forall>a. A.ide a \<longrightarrow> limit_cone J B (\<lambda>j. D (j, a)) (l a) (\<chi> a)"
shows "\<And>a. A.ide a \<Longrightarrow> L l \<chi> a = l a"
proof -
let ?L = "L l \<chi>"
let ?P = "P l \<chi>"
fix a
assume a: "A.ide a"
interpret \<chi>a: limit_cone J B \<open>\<lambda>j. D (j, a)\<close> \<open>l a\<close> \<open>\<chi> a\<close> using a assms by auto
have Pa: "?P a = (\<lambda>f. f \<in> B.hom (l a) (l a) \<and>
diagram.cones_map J B (\<lambda>j. D (j, a)) f (\<chi> a) = \<chi> a)"
using a vcomp_ide_dom \<chi>a.natural_transformation_axioms by simp
have "?P a (?L a)" using assms a L_arr [of l \<chi> a] by fastforce
moreover have "?P a (l a)"
proof -
have "?P a (l a) \<longleftrightarrow> l a \<in> B.hom (l a) (l a) \<and> \<chi>a.D.cones_map (l a) (\<chi> a) = \<chi> a"
using Pa by meson
thus ?thesis
using a \<chi>a.ide_apex \<chi>a.cone_axioms \<chi>a.D.cones_map_ide [of "\<chi> a" "l a"] by force
qed
moreover have "\<exists>!f. ?P a f"
using a Pa \<chi>a.is_universal \<chi>a.cone_axioms by force
ultimately show "?L a = l a" by blast
qed
lemma chosen_limits_induce_functor:
assumes "\<forall>a. A.ide a \<longrightarrow> limit_cone J B (\<lambda>j. D (j, a)) (l a) (\<chi> a)"
shows "functor A B (L l \<chi>)"
proof -
let ?L = "L l \<chi>"
let ?P = "\<lambda>a. \<lambda>f. \<guillemotleft>f : l (A.dom a) \<rightarrow>\<^sub>B l (A.cod a)\<guillemotright> \<and>
diagram.cones_map J B (\<lambda>j. D (j, A.cod a)) f (\<chi> (A.cod a))
= vertical_composite.map J B (\<chi> (A.dom a)) (\<lambda>j. D (j, a))"
interpret L: "functor" A B ?L
apply unfold_locales
using assms L_arr [of l] L_ide
apply auto[4]
proof -
fix a' a
assume 1: "A.arr (A a' a)"
have a: "A.arr a" using 1 by auto
have a': "\<guillemotleft>a' : A.cod a \<rightarrow>\<^sub>A A.cod a'\<guillemotright>" using 1 by auto
have a'a: "A.seq a' a" using 1 by auto
interpret \<chi>_dom_a: limit_cone J B \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>l (A.dom a)\<close> \<open>\<chi> (A.dom a)\<close>
using a assms by auto
interpret \<chi>_cod_a: limit_cone J B \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>l (A.cod a)\<close> \<open>\<chi> (A.cod a)\<close>
using a'a assms by auto
interpret \<chi>_dom_a'a: limit_cone J B \<open>\<lambda>j. D (j, A.dom (a' \<cdot>\<^sub>A a))\<close> \<open>l (A.dom (a' \<cdot>\<^sub>A a))\<close>
\<open>\<chi> (A.dom (a' \<cdot>\<^sub>A a))\<close>
using a'a assms by auto
interpret \<chi>_cod_a'a: limit_cone J B \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>l (A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>\<chi> (A.cod (a' \<cdot>\<^sub>A a))\<close>
using a'a assms by auto
interpret Da: natural_transformation J B
\<open>\<lambda>j. D (j, A.dom a)\<close> \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>\<lambda>j. D (j, a)\<close>
using a fixing_arr_gives_natural_transformation_2 by simp
interpret Da': natural_transformation J B
\<open>\<lambda>j. D (j, A.cod a)\<close> \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>\<lambda>j. D (j, a')\<close>
using a a'a fixing_arr_gives_natural_transformation_2 by fastforce
interpret Da'o\<chi>_cod_a: vertical_composite J B
\<chi>_cod_a.A.map \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>\<chi> (A.cod a)\<close> \<open>\<lambda>j. D (j, a')\<close>..
interpret Da'o\<chi>_cod_a: cone J B \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>l (A.cod a)\<close> Da'o\<chi>_cod_a.map
..
interpret Da'a: natural_transformation J B
\<open>\<lambda>j. D (j, A.dom (a' \<cdot>\<^sub>A a))\<close> \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>\<lambda>j. D (j, a' \<cdot>\<^sub>A a)\<close>
using a'a fixing_arr_gives_natural_transformation_2 [of "a' \<cdot>\<^sub>A a"] by auto
interpret Da'ao\<chi>_dom_a'a:
vertical_composite J B \<chi>_dom_a'a.A.map \<open>\<lambda>j. D (j, A.dom (a' \<cdot>\<^sub>A a))\<close>
\<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>\<chi> (A.dom (a' \<cdot>\<^sub>A a))\<close>
\<open>\<lambda>j. D (j, a' \<cdot>\<^sub>A a)\<close> ..
interpret Da'ao\<chi>_dom_a'a: cone J B \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>l (A.dom (a' \<cdot>\<^sub>A a))\<close> Da'ao\<chi>_dom_a'a.map ..
show "?L (a' \<cdot>\<^sub>A a) = ?L a' \<cdot>\<^sub>B ?L a"
proof -
have "?P (a' \<cdot>\<^sub>A a) (?L (a' \<cdot>\<^sub>A a))" using assms a'a L_arr [of l \<chi> "a' \<cdot>\<^sub>A a"] by fastforce
moreover have "?P (a' \<cdot>\<^sub>A a) (?L a' \<cdot>\<^sub>B ?L a)"
proof
have La: "\<guillemotleft>?L a : l (A.dom a) \<rightarrow>\<^sub>B l (A.cod a)\<guillemotright>"
using assms a L_arr by fast
moreover have La': "\<guillemotleft>?L a' : l (A.cod a) \<rightarrow>\<^sub>B l (A.cod a')\<guillemotright>"
using assms a a' L_arr [of l \<chi> a'] by auto
ultimately have seq: "B.seq (?L a') (?L a)" by (elim B.in_homE, auto)
thus La'_La: "\<guillemotleft>?L a' \<cdot>\<^sub>B ?L a : l (A.dom (a' \<cdot>\<^sub>A a)) \<rightarrow>\<^sub>B l (A.cod (a' \<cdot>\<^sub>A a))\<guillemotright>"
using a a' 1 La La' by (intro B.comp_in_homI, auto)
show "\<chi>_cod_a'a.D.cones_map (?L a' \<cdot>\<^sub>B ?L a) (\<chi> (A.cod (a' \<cdot>\<^sub>A a)))
= Da'ao\<chi>_dom_a'a.map"
proof -
have "\<chi>_cod_a'a.D.cones_map (?L a' \<cdot>\<^sub>B ?L a) (\<chi> (A.cod (a' \<cdot>\<^sub>A a)))
= (\<chi>_cod_a'a.D.cones_map (?L a) o \<chi>_cod_a'a.D.cones_map (?L a'))
(\<chi> (A.cod a'))"
proof -
have "\<chi>_cod_a'a.D.cones_map (?L a' \<cdot>\<^sub>B ?L a) (\<chi> (A.cod (a' \<cdot>\<^sub>A a))) =
restrict (\<chi>_cod_a'a.D.cones_map (?L a) \<circ> \<chi>_cod_a'a.D.cones_map (?L a'))
(\<chi>_cod_a'a.D.cones (B.cod (?L a')))
(\<chi> (A.cod (a' \<cdot>\<^sub>A a)))"
using seq \<chi>_cod_a'a.cone_axioms \<chi>_cod_a'a.D.cones_map_comp [of "?L a'" "?L a"]
by argo
also have "... = (\<chi>_cod_a'a.D.cones_map (?L a) o \<chi>_cod_a'a.D.cones_map (?L a'))
(\<chi> (A.cod a'))"
proof -
have "\<chi> (A.cod a') \<in> \<chi>_cod_a'a.D.cones (l (A.cod a'))"
using \<chi>_cod_a'a.cone_axioms a'a by simp
moreover have "B.cod (?L a') = l (A.cod a')"
using assms a' L_arr [of l] by auto
ultimately show ?thesis
using a' a'a by simp
qed
finally show ?thesis by blast
qed
also have "... = \<chi>_cod_a'a.D.cones_map (?L a)
(\<chi>_cod_a'a.D.cones_map (?L a') (\<chi> (A.cod a')))"
by simp
also have "... = \<chi>_cod_a'a.D.cones_map (?L a) Da'o\<chi>_cod_a.map"
proof -
have "?P a' (?L a')" using assms a' L_arr [of l \<chi> a'] by fast
moreover have
"?P a' = (\<lambda>f. f \<in> B.hom (l (A.cod a)) (l (A.cod a')) \<and>
\<chi>_cod_a'a.D.cones_map f (\<chi> (A.cod a')) = Da'o\<chi>_cod_a.map)"
using a'a by force
ultimately show ?thesis using a'a by force
qed
also have "... = vertical_composite.map J B
(\<chi>_cod_a.D.cones_map (?L a) (\<chi> (A.cod a)))
(\<lambda>j. D (j, a'))"
using assms \<chi>_cod_a.D.diagram_axioms \<chi>_cod_a'a.D.diagram_axioms
Da'.natural_transformation_axioms \<chi>_cod_a.cone_axioms La
cones_map_vcomp [of J B "\<lambda>j. D (j, A.cod a)" "\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))"
"\<lambda>j. D (j, a')" "l (A.cod a)" "\<chi> (A.cod a)"
"?L a" "l (A.dom a)"]
by blast
also have "... = vertical_composite.map J B
(vertical_composite.map J B (\<chi> (A.dom a)) (\<lambda>j. D (j, a)))
(\<lambda>j. D (j, a'))"
using assms a L_arr by presburger
also have "... = vertical_composite.map J B (\<chi> (A.dom a))
(vertical_composite.map J B (\<lambda>j. D (j, a)) (\<lambda>j. D (j, a')))"
using a'a Da.natural_transformation_axioms Da'.natural_transformation_axioms
\<chi>_dom_a.natural_transformation_axioms vcomp_assoc
by auto
also have
"... = vertical_composite.map J B (\<chi> (A.dom (a' \<cdot>\<^sub>A a))) (\<lambda>j. D (j, a' \<cdot>\<^sub>A a))"
using a'a preserves_comp_2 by simp
finally show ?thesis by auto
qed
qed
moreover have "\<exists>!f. ?P (a' \<cdot>\<^sub>A a) f"
using \<chi>_cod_a'a.is_universal
[of "l (A.dom (a' \<cdot>\<^sub>A a))"
"vertical_composite.map J B (\<chi> (A.dom (a' \<cdot>\<^sub>A a))) (\<lambda>j. D (j, a' \<cdot>\<^sub>A a))"]
Da'ao\<chi>_dom_a'a.cone_axioms
by fast
ultimately show ?thesis by blast
qed
qed
show ?thesis ..
qed
end
locale diagram_in_functor_category =
A: category A +
B: category B +
A_B: functor_category A B +
diagram J A_B.comp D
for A :: "'a comp" (infixr "\<cdot>\<^sub>A" 55)
and B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and D :: "'j \<Rightarrow> ('a, 'b) functor_category.arr"
begin
interpretation JxA: product_category J A ..
interpretation A_BxA: product_category A_B.comp A ..
interpretation E: evaluation_functor A B ..
interpretation Curry: currying J A B ..
notation JxA.comp (infixr "\<cdot>\<^sub>J\<^sub>x\<^sub>A" 55)
notation JxA.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J\<^sub>x\<^sub>A _\<guillemotright>")
text\<open>
Evaluation of a functor or natural transformation from @{term[source=true] J}
to \<open>[A, B]\<close> at an arrow @{term a} of @{term[source=true] A}.
\<close>
abbreviation at
where "at a \<tau> \<equiv> \<lambda>j. Curry.uncurry \<tau> (j, a)"
lemma at_simp:
assumes "A.arr a" and "J.arr j" and "A_B.arr (\<tau> j)"
shows "at a \<tau> j = A_B.Map (\<tau> j) a"
using assms Curry.uncurry_def E.map_simp by simp
lemma functor_at_ide_is_functor:
assumes "functor J A_B.comp F" and "A.ide a"
shows "functor J B (at a F)"
proof -
interpret uncurry_F: "functor" JxA.comp B \<open>Curry.uncurry F\<close>
using assms(1) Curry.uncurry_preserves_functors by simp
interpret uncurry_F: binary_functor J A B \<open>Curry.uncurry F\<close> ..
show ?thesis using assms(2) uncurry_F.fixing_ide_gives_functor_2 by simp
qed
lemma functor_at_arr_is_transformation:
assumes "functor J A_B.comp F" and "A.arr a"
shows "natural_transformation J B (at (A.dom a) F) (at (A.cod a) F) (at a F)"
proof -
interpret uncurry_F: "functor" JxA.comp B \<open>Curry.uncurry F\<close>
using assms(1) Curry.uncurry_preserves_functors by simp
interpret uncurry_F: binary_functor J A B \<open>Curry.uncurry F\<close> ..
show ?thesis
using assms(2) uncurry_F.fixing_arr_gives_natural_transformation_2 by simp
qed
lemma transformation_at_ide_is_transformation:
assumes "natural_transformation J A_B.comp F G \<tau>" and "A.ide a"
shows "natural_transformation J B (at a F) (at a G) (at a \<tau>)"
proof -
interpret \<tau>: natural_transformation J A_B.comp F G \<tau> using assms(1) by auto
interpret uncurry_F: "functor" JxA.comp B \<open>Curry.uncurry F\<close>
using Curry.uncurry_preserves_functors \<tau>.F.functor_axioms by simp
interpret uncurry_f: binary_functor J A B \<open>Curry.uncurry F\<close> ..
interpret uncurry_G: "functor" JxA.comp B \<open>Curry.uncurry G\<close>
using Curry.uncurry_preserves_functors \<tau>.G.functor_axioms by simp
interpret uncurry_G: binary_functor J A B \<open>Curry.uncurry G\<close> ..
interpret uncurry_\<tau>: natural_transformation
JxA.comp B \<open>Curry.uncurry F\<close> \<open>Curry.uncurry G\<close> \<open>Curry.uncurry \<tau>\<close>
using Curry.uncurry_preserves_transformations \<tau>.natural_transformation_axioms
by simp
interpret uncurry_\<tau>: binary_functor_transformation J A B
\<open>Curry.uncurry F\<close> \<open>Curry.uncurry G\<close> \<open>Curry.uncurry \<tau>\<close> ..
show ?thesis
using assms(2) uncurry_\<tau>.fixing_ide_gives_natural_transformation_2 by simp
qed
lemma constant_at_ide_is_constant:
assumes "cone x \<chi>" and a: "A.ide a"
shows "at a (constant_functor.map J A_B.comp x) =
constant_functor.map J B (A_B.Map x a)"
proof -
interpret \<chi>: cone J A_B.comp D x \<chi> using assms(1) by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
interpret Fun_x: "functor" A B \<open>A_B.Map x\<close>
using x A_B.ide_char by simp
interpret Da: "functor" J B \<open>at a D\<close>
using a functor_at_ide_is_functor functor_axioms by blast
interpret Da: diagram J B \<open>at a D\<close> ..
interpret Xa: constant_functor J B \<open>A_B.Map x a\<close>
using a Fun_x.preserves_ide by unfold_locales simp
show "at a \<chi>.A.map = Xa.map"
using a x Curry.uncurry_def E.map_def Xa.is_extensional by auto
qed
lemma at_ide_is_diagram:
assumes a: "A.ide a"
shows "diagram J B (at a D)"
proof -
interpret Da: "functor" J B "at a D"
using a functor_at_ide_is_functor functor_axioms by simp
show ?thesis ..
qed
lemma cone_at_ide_is_cone:
assumes "cone x \<chi>" and a: "A.ide a"
shows "diagram.cone J B (at a D) (A_B.Map x a) (at a \<chi>)"
proof -
interpret \<chi>: cone J A_B.comp D x \<chi> using assms(1) by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
interpret Fun_x: "functor" A B \<open>A_B.Map x\<close>
using x A_B.ide_char by simp
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret Xa: constant_functor J B \<open>A_B.Map x a\<close>
using a by (unfold_locales, simp)
interpret \<chi>a: natural_transformation J B Xa.map \<open>at a D\<close> \<open>at a \<chi>\<close>
using assms(1) x a transformation_at_ide_is_transformation \<chi>.natural_transformation_axioms
constant_at_ide_is_constant
by fastforce
interpret \<chi>a: cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close> ..
show cone_\<chi>a: "Da.cone (A_B.Map x a) (at a \<chi>)" ..
qed
lemma at_preserves_comp:
assumes "A.seq a' a"
shows "at (A a' a) D = vertical_composite.map J B (at a D) (at a' D)"
proof -
interpret Da: natural_transformation J B \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close> \<open>at a D\<close>
using assms functor_at_arr_is_transformation functor_axioms by blast
interpret Da': natural_transformation J B \<open>at (A.cod a) D\<close> \<open>at (A.cod a') D\<close> \<open>at a' D\<close>
using assms functor_at_arr_is_transformation [of D a'] functor_axioms by fastforce
interpret Da'oDa: vertical_composite J B
\<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close> \<open>at (A.cod a') D\<close>
\<open>at a D\<close> \<open>at a' D\<close> ..
interpret Da'a: natural_transformation J B \<open>at (A.dom a) D\<close> \<open>at (A.cod a') D\<close>
\<open>at (a' \<cdot>\<^sub>A a) D\<close>
using assms functor_at_arr_is_transformation [of D "a' \<cdot>\<^sub>A a"] functor_axioms by simp
show "at (a' \<cdot>\<^sub>A a) D = Da'oDa.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J B (at (A.dom a) D) (at (A.cod a') D) Da'oDa.map" ..
show "natural_transformation J B (at (A.dom a) D) (at (A.cod a') D) (at (a' \<cdot>\<^sub>A a) D)" ..
show "\<And>j. J.ide j \<Longrightarrow> at (a' \<cdot>\<^sub>A a) D j = Da'oDa.map j"
proof -
fix j
assume j: "J.ide j"
interpret Dj: "functor" A B \<open>A_B.Map (D j)\<close>
using j preserves_ide A_B.ide_char by simp
show "at (a' \<cdot>\<^sub>A a) D j = Da'oDa.map j"
using assms j Dj.preserves_comp at_simp Da'oDa.map_simp_ide by auto
qed
qed
qed
lemma cones_map_pointwise:
assumes "cone x \<chi>" and "cone x' \<chi>'"
and f: "f \<in> A_B.hom x' x"
shows "cones_map f \<chi> = \<chi>' \<longleftrightarrow>
(\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>')"
proof
interpret \<chi>: cone J A_B.comp D x \<chi> using assms(1) by auto
interpret \<chi>': cone J A_B.comp D x' \<chi>' using assms(2) by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
have x': "A_B.ide x'" using \<chi>'.ide_apex by auto
interpret \<chi>f: cone J A_B.comp D x' \<open>cones_map f \<chi>\<close>
using x' f assms(1) cones_map_mapsto by blast
interpret Fun_x: "functor" A B \<open>A_B.Map x\<close> using x A_B.ide_char by simp
interpret Fun_x': "functor" A B \<open>A_B.Map x'\<close> using x' A_B.ide_char by simp
show "cones_map f \<chi> = \<chi>' \<Longrightarrow>
(\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>')"
proof -
assume \<chi>': "cones_map f \<chi> = \<chi>'"
have "\<And>a. A.ide a \<Longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret \<chi>a: cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close>
using a assms(1) cone_at_ide_is_cone by simp
interpret \<chi>'a: cone J B \<open>at a D\<close> \<open>A_B.Map x' a\<close> \<open>at a \<chi>'\<close>
using a assms(2) cone_at_ide_is_cone by simp
have 1: "\<guillemotleft>A_B.Map f a : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright>"
using f a A_B.arr_char A_B.Map_cod A_B.Map_dom mem_Collect_eq
natural_transformation.preserves_hom A.ide_in_hom
by (metis (no_types, lifting) A_B.in_homE)
interpret \<chi>fa: cone J B \<open>at a D\<close> \<open>A_B.Map x' a\<close>
\<open>Da.cones_map (A_B.Map f a) (at a \<chi>)\<close>
using 1 \<chi>a.cone_axioms Da.cones_map_mapsto by force
show "Da.cones_map (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> Da.cones_map (A_B.Map f a) (at a \<chi>) j = at a \<chi>' j"
using \<chi>'a.is_extensional \<chi>fa.is_extensional [of j] by simp
moreover have "J.arr j \<Longrightarrow> Da.cones_map (A_B.Map f a) (at a \<chi>) j = at a \<chi>' j"
using a f 1 \<chi>.cone_axioms \<chi>a.cone_axioms at_simp
apply simp
apply (elim A_B.in_homE B.in_homE, auto)
using \<chi>' \<chi>.A.map_simp A_B.Map_comp [of "\<chi> j" f a a] by auto
ultimately show "Da.cones_map (A_B.Map f a) (at a \<chi>) j = at a \<chi>' j" by blast
qed
qed
thus "\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
by simp
qed
show "\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'
\<Longrightarrow> cones_map f \<chi> = \<chi>'"
proof -
assume A:
"\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
show "cones_map f \<chi> = \<chi>'"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J A_B.comp \<chi>'.A.map D (cones_map f \<chi>)" ..
show "natural_transformation J A_B.comp \<chi>'.A.map D \<chi>'" ..
show "\<And>j. J.ide j \<Longrightarrow> cones_map f \<chi> j = \<chi>' j"
proof (intro A_B.arr_eqI)
fix j
assume j: "J.ide j"
show 1: "A_B.arr (cones_map f \<chi> j)"
using j \<chi>f.preserves_reflects_arr by simp
show "A_B.arr (\<chi>' j)" using j by auto
have Dom_\<chi>f_j: "A_B.Dom (cones_map f \<chi> j) = A_B.Map x'"
using x' j 1 A_B.Map_dom \<chi>'.A.map_simp \<chi>f.preserves_dom J.ide_in_hom
by (metis (no_types, lifting) J.ideD(2) \<chi>f.preserves_reflects_arr)
also have Dom_\<chi>'_j: "... = A_B.Dom (\<chi>' j)"
using x' j A_B.Map_dom [of "\<chi>' j"] \<chi>'.preserves_hom \<chi>'.A.map_simp by simp
finally show "A_B.Dom (cones_map f \<chi> j) = A_B.Dom (\<chi>' j)" by auto
have Cod_\<chi>f_j: "A_B.Cod (cones_map f \<chi> j) = A_B.Map (D (J.cod j))"
using j A_B.Map_cod A_B.cod_char J.ide_in_hom \<chi>f.preserves_hom
by (metis (no_types, lifting) "1" J.ideD(1) \<chi>f.preserves_cod)
also have Cod_\<chi>'_j: "... = A_B.Cod (\<chi>' j)"
using j A_B.Map_cod [of "\<chi>' j"] \<chi>'.preserves_hom by simp
finally show "A_B.Cod (cones_map f \<chi> j) = A_B.Cod (\<chi>' j)" by auto
show "A_B.Map (cones_map f \<chi> j) = A_B.Map (\<chi>' j)"
proof (intro NaturalTransformation.eqI)
interpret \<chi>fj: natural_transformation A B \<open>A_B.Map x'\<close> \<open>A_B.Map (D (J.cod j))\<close>
\<open>A_B.Map (cones_map f \<chi> j)\<close>
using j \<chi>f.preserves_reflects_arr A_B.arr_char [of "cones_map f \<chi> j"]
Dom_\<chi>f_j Cod_\<chi>f_j
by simp
show "natural_transformation A B (A_B.Map x') (A_B.Map (D (J.cod j)))
(A_B.Map (cones_map f \<chi> j))" ..
interpret \<chi>'j: natural_transformation A B \<open>A_B.Map x'\<close> \<open>A_B.Map (D (J.cod j))\<close>
\<open>A_B.Map (\<chi>' j)\<close>
using j A_B.arr_char [of "\<chi>' j"] Dom_\<chi>'_j Cod_\<chi>'_j by simp
show "natural_transformation A B (A_B.Map x') (A_B.Map (D (J.cod j)))
(A_B.Map (\<chi>' j))" ..
show "\<And>a. A.ide a \<Longrightarrow> A_B.Map (cones_map f \<chi> j) a = A_B.Map (\<chi>' j) a"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
have cone_\<chi>a: "Da.cone (A_B.Map x a) (at a \<chi>)"
using a assms(1) cone_at_ide_is_cone by simp
interpret \<chi>a: cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close>
using cone_\<chi>a by auto
interpret Fun_f: natural_transformation A B \<open>A_B.Dom f\<close> \<open>A_B.Cod f\<close>
\<open>A_B.Map f\<close>
using f A_B.arr_char by fast
have fa: "A_B.Map f a \<in> B.hom (A_B.Map x' a) (A_B.Map x a)"
using a f Fun_f.preserves_hom A.ide_in_hom by auto
have "A_B.Map (cones_map f \<chi> j) a = Da.cones_map (A_B.Map f a) (at a \<chi>) j"
proof -
have "A_B.Map (cones_map f \<chi> j) a = A_B.Map (A_B.comp (\<chi> j) f) a"
using assms(1) f \<chi>.is_extensional by auto
also have "... = B (A_B.Map (\<chi> j) a) (A_B.Map f a)"
using f j a \<chi>.preserves_hom A.ide_in_hom J.ide_in_hom A_B.Map_comp
\<chi>.A.map_simp
by (metis (no_types, lifting) A.comp_ide_self A.ideD(1) A_B.seqI'
J.ideD(1) mem_Collect_eq)
also have "... = Da.cones_map (A_B.Map f a) (at a \<chi>) j"
using j a cone_\<chi>a fa Curry.uncurry_def E.map_simp by auto
finally show ?thesis by auto
qed
also have "... = at a \<chi>' j" using j a A by simp
also have "... = A_B.Map (\<chi>' j) a"
using j Curry.uncurry_def E.map_simp \<chi>'j.is_extensional by simp
finally show "A_B.Map (cones_map f \<chi> j) a = A_B.Map (\<chi>' j) a" by auto
qed
qed
qed
qed
qed
qed
text\<open>
If @{term \<chi>} is a cone with apex @{term a} over @{term D}, then @{term \<chi>}
is a limit cone if, for each object @{term x} of @{term X}, the cone obtained
by evaluating @{term \<chi>} at @{term x} is a limit cone with apex @{term "A_B.Map a x"}
for the diagram in @{term C} obtained by evaluating @{term D} at @{term x}.
\<close>
lemma cone_is_limit_if_pointwise_limit:
assumes cone_\<chi>: "cone x \<chi>"
and "\<forall>a. A.ide a \<longrightarrow> diagram.limit_cone J B (at a D) (A_B.Map x a) (at a \<chi>)"
shows "limit_cone x \<chi>"
proof -
interpret \<chi>: cone J A_B.comp D x \<chi> using assms by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
show "limit_cone x \<chi>"
proof
fix x' \<chi>'
assume cone_\<chi>': "cone x' \<chi>'"
interpret \<chi>': cone J A_B.comp D x' \<chi>' using cone_\<chi>' by auto
have x': "A_B.ide x'" using \<chi>'.ide_apex by auto
text\<open>
The universality of the limit cone \<open>at a \<chi>\<close> yields, for each object
\<open>a\<close> of \<open>A\<close>, a unique arrow \<open>fa\<close> that transforms
\<open>at a \<chi>\<close> to \<open>at a \<chi>'\<close>.
\<close>
have EU: "\<And>a. A.ide a \<Longrightarrow>
\<exists>!fa. fa \<in> B.hom (A_B.Map x' a) (A_B.Map x a) \<and>
diagram.cones_map J B (at a D) fa (at a \<chi>) = at a \<chi>'"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret \<chi>a: limit_cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close>
using assms(2) a by auto
interpret \<chi>'a: cone J B \<open>at a D\<close> \<open>A_B.Map x' a\<close> \<open>at a \<chi>'\<close>
using a cone_\<chi>' cone_at_ide_is_cone by auto
have "Da.cone (A_B.Map x' a) (at a \<chi>')" ..
thus "\<exists>!fa. fa \<in> B.hom (A_B.Map x' a) (A_B.Map x a) \<and>
Da.cones_map fa (at a \<chi>) = at a \<chi>'"
using \<chi>a.is_universal by simp
qed
text\<open>
Our objective is to show the existence of a unique arrow \<open>f\<close> that transforms
\<open>\<chi>\<close> into \<open>\<chi>'\<close>. We obtain \<open>f\<close> by bundling the arrows \<open>fa\<close>
of \<open>C\<close> and proving that this yields a natural transformation from \<open>X\<close>
to \<open>C\<close>, hence an arrow of \<open>[X, C]\<close>.
\<close>
show "\<exists>!f. \<guillemotleft>f : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map f \<chi> = \<chi>'"
proof
let ?P = "\<lambda>a fa. \<guillemotleft>fa : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright> \<and>
diagram.cones_map J B (at a D) fa (at a \<chi>) = at a \<chi>'"
have AaPa: "\<And>a. A.ide a \<Longrightarrow> ?P a (THE fa. ?P a fa)"
proof -
fix a
assume a: "A.ide a"
have "\<exists>!fa. ?P a fa" using a EU by simp
thus "?P a (THE fa. ?P a fa)" using a theI' [of "?P a"] by fastforce
qed
have AaPa_in_hom:
"\<And>a. A.ide a \<Longrightarrow> \<guillemotleft>THE fa. ?P a fa : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright>"
using AaPa by blast
have AaPa_map:
"\<And>a. A.ide a \<Longrightarrow>
diagram.cones_map J B (at a D) (THE fa. ?P a fa) (at a \<chi>) = at a \<chi>'"
using AaPa by blast
let ?Fun_f = "\<lambda>a. if A.ide a then (THE fa. ?P a fa) else B.null"
interpret Fun_x: "functor" A B \<open>\<lambda>a. A_B.Map x a\<close>
using x A_B.ide_char by simp
interpret Fun_x': "functor" A B \<open>\<lambda>a. A_B.Map x' a\<close>
using x' A_B.ide_char by simp
text\<open>
The arrows \<open>Fun_f a\<close> are the components of a natural transformation.
It is more work to verify the naturality than it seems like it ought to be.
\<close>
interpret \<phi>: transformation_by_components A B
\<open>\<lambda>a. A_B.Map x' a\<close> \<open>\<lambda>a. A_B.Map x a\<close> ?Fun_f
proof
fix a
assume a: "A.ide a"
show "\<guillemotleft>?Fun_f a : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright>" using a AaPa by simp
next
fix a
assume a: "A.arr a"
text\<open>
\newcommand\xdom{\mathop{\rm dom}}
\newcommand\xcod{\mathop{\rm cod}}
$$\xymatrix{
{x_{\xdom a}} \drtwocell\omit{\omit(A)} \ar[d]_{\chi_{\xdom a}} \ar[r]^{x_a} & {x_{\xcod a}}
\ar[d]^{\chi_{\xcod a}} \\
{D_{\xdom a}} \ar[r]^{D_a} & {D_{\xcod a}} \\
{x'_{\xdom a}} \urtwocell\omit{\omit(B)} \ar@/^5em/[uu]^{f_{\xdom a}}_{\hspace{1em}(C)} \ar[u]^{\chi'_{\xdom a}}
\ar[r]_{x'_a} & {x'_{\xcod a}} \ar[u]_{x'_{\xcod a}} \ar@/_5em/[uu]_{f_{\xcod a}}
}$$
\<close>
let ?x_dom_a = "A_B.Map x (A.dom a)"
let ?x_cod_a = "A_B.Map x (A.cod a)"
let ?x_a = "A_B.Map x a"
have x_a: "\<guillemotleft>?x_a : ?x_dom_a \<rightarrow>\<^sub>B ?x_cod_a\<guillemotright>"
using a x A_B.ide_char by auto
let ?x'_dom_a = "A_B.Map x' (A.dom a)"
let ?x'_cod_a = "A_B.Map x' (A.cod a)"
let ?x'_a = "A_B.Map x' a"
have x'_a: "\<guillemotleft>?x'_a : ?x'_dom_a \<rightarrow>\<^sub>B ?x'_cod_a\<guillemotright>"
using a x' A_B.ide_char by auto
let ?f_dom_a = "?Fun_f (A.dom a)"
let ?f_cod_a = "?Fun_f (A.cod a)"
have f_dom_a: "\<guillemotleft>?f_dom_a : ?x'_dom_a \<rightarrow>\<^sub>B ?x_dom_a\<guillemotright>" using a AaPa by simp
have f_cod_a: "\<guillemotleft>?f_cod_a : ?x'_cod_a \<rightarrow>\<^sub>B ?x_cod_a\<guillemotright>" using a AaPa by simp
interpret D_dom_a: diagram J B \<open>at (A.dom a) D\<close> using a at_ide_is_diagram by simp
interpret D_cod_a: diagram J B \<open>at (A.cod a) D\<close> using a at_ide_is_diagram by simp
interpret Da: natural_transformation J B \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close> \<open>at a D\<close>
using a functor_axioms functor_at_arr_is_transformation by simp
interpret \<chi>_dom_a: limit_cone J B \<open>at (A.dom a) D\<close> \<open>A_B.Map x (A.dom a)\<close>
\<open>at (A.dom a) \<chi>\<close>
using assms(2) a by auto
interpret \<chi>_cod_a: limit_cone J B \<open>at (A.cod a) D\<close> \<open>A_B.Map x (A.cod a)\<close>
\<open>at (A.cod a) \<chi>\<close>
using assms(2) a by auto
interpret \<chi>'_dom_a: cone J B \<open>at (A.dom a) D\<close> \<open>A_B.Map x' (A.dom a)\<close>
\<open>at (A.dom a) \<chi>'\<close>
using a cone_\<chi>' cone_at_ide_is_cone by auto
interpret \<chi>'_cod_a: cone J B \<open>at (A.cod a) D\<close> \<open>A_B.Map x' (A.cod a)\<close>
\<open>at (A.cod a) \<chi>'\<close>
using a cone_\<chi>' cone_at_ide_is_cone by auto
text\<open>
Now construct cones with apexes \<open>x_dom_a\<close> and \<open>x'_dom_a\<close>
over @{term "at (A.cod a) D"} by forming the vertical composites of
@{term "at (A.dom a) \<chi>"} and @{term "at (A.cod a) \<chi>'"} with the natural
transformation @{term "at a D"}.
\<close>
interpret Dao\<chi>_dom_a: vertical_composite J B
\<chi>_dom_a.A.map \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close>
\<open>at (A.dom a) \<chi>\<close> \<open>at a D\<close> ..
interpret Dao\<chi>_dom_a: cone J B \<open>at (A.cod a) D\<close> ?x_dom_a Dao\<chi>_dom_a.map
using \<chi>_dom_a.cone_axioms Da.natural_transformation_axioms vcomp_transformation_cone
by metis
interpret Dao\<chi>'_dom_a: vertical_composite J B
\<chi>'_dom_a.A.map \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close>
\<open>at (A.dom a) \<chi>'\<close> \<open>at a D\<close> ..
interpret Dao\<chi>'_dom_a: cone J B \<open>at (A.cod a) D\<close> ?x'_dom_a Dao\<chi>'_dom_a.map
using \<chi>'_dom_a.cone_axioms Da.natural_transformation_axioms vcomp_transformation_cone
by metis
have Dao\<chi>_dom_a: "D_cod_a.cone ?x_dom_a Dao\<chi>_dom_a.map" ..
have Dao\<chi>'_dom_a: "D_cod_a.cone ?x'_dom_a Dao\<chi>'_dom_a.map" ..
text\<open>
These cones are also obtained by transforming the cones @{term "at (A.cod a) \<chi>"}
and @{term "at (A.cod a) \<chi>'"} by \<open>x_a\<close> and \<open>x'_a\<close>, respectively.
\<close>
have A: "Dao\<chi>_dom_a.map = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>)"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> Dao\<chi>_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
using Dao\<chi>_dom_a.is_extensional \<chi>_cod_a.cone_axioms x_a by force
moreover have
"J.arr j \<Longrightarrow> Dao\<chi>_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
proof -
assume j: "J.arr j"
have "Dao\<chi>_dom_a.map j = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi> (J.dom j)"
using j Dao\<chi>_dom_a.map_simp_2 by simp
also have "... = A_B.Map (D j) a \<cdot>\<^sub>B A_B.Map (\<chi> (J.dom j)) (A.dom a)"
using a j at_simp by simp
also have "... = A_B.Map (A_B.comp (D j) (\<chi> (J.dom j))) a"
using a j A_B.Map_comp
by (metis (no_types, lifting) A.comp_arr_dom \<chi>.is_natural_1
\<chi>.preserves_reflects_arr)
also have "... = A_B.Map (A_B.comp (\<chi> (J.cod j)) (\<chi>.A.map j)) a"
using a j \<chi>.naturality by simp
also have "... = A_B.Map (\<chi> (J.cod j)) (A.cod a) \<cdot>\<^sub>B A_B.Map x a"
using a j x A_B.Map_comp
by (metis (no_types, lifting) A.comp_cod_arr \<chi>.A.map_simp \<chi>.is_natural_2
\<chi>.preserves_reflects_arr)
also have "... = at (A.cod a) \<chi> (J.cod j) \<cdot>\<^sub>B A_B.Map x a"
using a j at_simp by simp
also have "... = at (A.cod a) \<chi> j \<cdot>\<^sub>B A_B.Map x a"
using a j \<chi>_cod_a.is_natural_2 \<chi>_cod_a.A.map_simp
by (metis J.arr_cod_iff_arr J.cod_cod)
also have "... = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
using a j x \<chi>_cod_a.cone_axioms preserves_cod by simp
finally show ?thesis by blast
qed
ultimately show "Dao\<chi>_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
by blast
qed
have B: "Dao\<chi>'_dom_a.map = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>')"
proof
fix j
have "\<not>J.arr j \<Longrightarrow>
Dao\<chi>'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
using Dao\<chi>'_dom_a.is_extensional \<chi>'_cod_a.cone_axioms x'_a by force
moreover have
"J.arr j \<Longrightarrow> Dao\<chi>'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
proof -
assume j: "J.arr j"
have "Dao\<chi>'_dom_a.map j = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi>' (J.dom j)"
using j Dao\<chi>'_dom_a.map_simp_2 by simp
also have "... = A_B.Map (D j) a \<cdot>\<^sub>B A_B.Map (\<chi>' (J.dom j)) (A.dom a)"
using a j at_simp by simp
also have "... = A_B.Map (A_B.comp (D j) (\<chi>' (J.dom j))) a"
using a j A_B.Map_comp
by (metis (no_types, lifting) A.comp_arr_dom \<chi>'.is_natural_1
\<chi>'.preserves_reflects_arr)
also have "... = A_B.Map (A_B.comp (\<chi>' (J.cod j)) (\<chi>'.A.map j)) a"
using a j \<chi>'.naturality by simp
also have "... = A_B.Map (\<chi>' (J.cod j)) (A.cod a) \<cdot>\<^sub>B A_B.Map x' a"
using a j x' A_B.Map_comp
by (metis (no_types, lifting) A.comp_cod_arr \<chi>'.A.map_simp \<chi>'.is_natural_2
\<chi>'.preserves_reflects_arr)
also have "... = at (A.cod a) \<chi>' (J.cod j) \<cdot>\<^sub>B A_B.Map x' a"
using a j at_simp by simp
also have "... = at (A.cod a) \<chi>' j \<cdot>\<^sub>B A_B.Map x' a"
using a j \<chi>'_cod_a.is_natural_2 \<chi>'_cod_a.A.map_simp
by (metis J.arr_cod_iff_arr J.cod_cod)
also have "... = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
using a j x' \<chi>'_cod_a.cone_axioms preserves_cod by simp
finally show ?thesis by blast
qed
ultimately show
"Dao\<chi>'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
by blast
qed
text\<open>
Next, we show that \<open>f_dom_a\<close>, which is the unique arrow that transforms
\<open>\<chi>_dom_a\<close> into \<open>\<chi>'_dom_a\<close>, is also the unique arrow that transforms
\<open>Dao\<chi>_dom_a\<close> into \<open>Dao\<chi>'_dom_a\<close>.
\<close>
have C: "D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map = Dao\<chi>'_dom_a.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation
J B \<chi>'_dom_a.A.map (at (A.cod a) D) Dao\<chi>'_dom_a.map" ..
show "natural_transformation J B \<chi>'_dom_a.A.map (at (A.cod a) D)
(D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map)"
proof -
interpret \<kappa>: cone J B \<open>at (A.cod a) D\<close> ?x'_dom_a
\<open>D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map\<close>
proof -
have "\<And>b b' f. \<lbrakk> f \<in> B.hom b' b; D_cod_a.cone b Dao\<chi>_dom_a.map \<rbrakk>
\<Longrightarrow> D_cod_a.cone b' (D_cod_a.cones_map f Dao\<chi>_dom_a.map)"
using D_cod_a.cones_map_mapsto by blast
moreover have "D_cod_a.cone ?x_dom_a Dao\<chi>_dom_a.map" ..
ultimately show "D_cod_a.cone ?x'_dom_a
(D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map)"
using f_dom_a by simp
qed
show ?thesis ..
qed
show "\<And>j. J.ide j \<Longrightarrow>
D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map j = Dao\<chi>'_dom_a.map j"
proof -
fix j
assume j: "J.ide j"
have "D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map j =
Dao\<chi>_dom_a.map j \<cdot>\<^sub>B ?f_dom_a"
using j f_dom_a Dao\<chi>_dom_a.cone_axioms
by (elim B.in_homE, auto)
also have "... = (at a D j \<cdot>\<^sub>B at (A.dom a) \<chi> j) \<cdot>\<^sub>B ?f_dom_a"
using j Dao\<chi>_dom_a.map_simp_ide by simp
also have "... = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi> j \<cdot>\<^sub>B ?f_dom_a"
using B.comp_assoc by simp
also have "... = at a D j \<cdot>\<^sub>B D_dom_a.cones_map ?f_dom_a (at (A.dom a) \<chi>) j"
using j \<chi>_dom_a.cone_axioms f_dom_a
by (elim B.in_homE, auto)
also have "... = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi>' j"
using a AaPa A.ide_dom by presburger
also have "... = Dao\<chi>'_dom_a.map j"
using j Dao\<chi>'_dom_a.map_simp_ide by simp
finally show
"D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map j = Dao\<chi>'_dom_a.map j"
by auto
qed
qed
text\<open>
Naturality amounts to showing that \<open>C f_cod_a x'_a = C x_a f_dom_a\<close>.
To do this, we show that both arrows transform @{term "at (A.cod a) \<chi>"}
into \<open>Dao\<chi>'_cod_a\<close>, thus they are equal by the universality of
@{term "at (A.cod a) \<chi>"}.
\<close>
have "\<exists>!fa. \<guillemotleft>fa : ?x'_dom_a \<rightarrow>\<^sub>B ?x_cod_a\<guillemotright> \<and>
D_cod_a.cones_map fa (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
using Dao\<chi>'_dom_a.cone_axioms a \<chi>_cod_a.is_universal [of ?x'_dom_a Dao\<chi>'_dom_a.map]
by fast
moreover have
"?f_cod_a \<cdot>\<^sub>B ?x'_a \<in> B.hom ?x'_dom_a ?x_cod_a \<and>
D_cod_a.cones_map (?f_cod_a \<cdot>\<^sub>B ?x'_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof
show "?f_cod_a \<cdot>\<^sub>B ?x'_a \<in> B.hom ?x'_dom_a ?x_cod_a"
using f_cod_a x'_a by blast
show "D_cod_a.cones_map (?f_cod_a \<cdot>\<^sub>B ?x'_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof -
have "D_cod_a.cones_map (?f_cod_a \<cdot>\<^sub>B ?x'_a) (at (A.cod a) \<chi>)
= restrict (D_cod_a.cones_map ?x'_a o D_cod_a.cones_map ?f_cod_a)
(D_cod_a.cones (?x_cod_a))
(at (A.cod a) \<chi>)"
using x'_a D_cod_a.cones_map_comp [of ?f_cod_a ?x'_a] f_cod_a
by (elim B.in_homE, auto)
also have "... = D_cod_a.cones_map ?x'_a
(D_cod_a.cones_map ?f_cod_a (at (A.cod a) \<chi>))"
using \<chi>_cod_a.cone_axioms by simp
also have "... = Dao\<chi>'_dom_a.map"
using a B AaPa_map A.ide_cod by presburger
finally show ?thesis by auto
qed
qed
moreover have
"?x_a \<cdot>\<^sub>B ?f_dom_a \<in> B.hom ?x'_dom_a ?x_cod_a \<and>
D_cod_a.cones_map (?x_a \<cdot>\<^sub>B ?f_dom_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof
show "?x_a \<cdot>\<^sub>B ?f_dom_a \<in> B.hom ?x'_dom_a ?x_cod_a"
using f_dom_a x_a by blast
show "D_cod_a.cones_map (?x_a \<cdot>\<^sub>B ?f_dom_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof -
have
"D_cod_a.cones (B.cod (A_B.Map x a)) = D_cod_a.cones (A_B.Map x (A.cod a))"
using a x by simp
moreover have "B.seq ?x_a ?f_dom_a"
using f_dom_a x_a by (elim B.in_homE, auto)
ultimately have
"D_cod_a.cones_map (?x_a \<cdot>\<^sub>B ?f_dom_a) (at (A.cod a) \<chi>)
= restrict (D_cod_a.cones_map ?f_dom_a o D_cod_a.cones_map ?x_a)
(D_cod_a.cones (?x_cod_a))
(at (A.cod a) \<chi>)"
using D_cod_a.cones_map_comp [of ?x_a ?f_dom_a] x_a by argo
also have "... = D_cod_a.cones_map ?f_dom_a
(D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>))"
using \<chi>_cod_a.cone_axioms by simp
also have "... = Dao\<chi>'_dom_a.map"
using A C a AaPa by argo
finally show ?thesis by blast
qed
qed
ultimately show "?f_cod_a \<cdot>\<^sub>B ?x'_a = ?x_a \<cdot>\<^sub>B ?f_dom_a"
using a \<chi>_cod_a.is_universal by blast
qed
text\<open>
The arrow from @{term x'} to @{term x} in \<open>[A, B]\<close> determined by
the natural transformation \<open>\<phi>\<close> transforms @{term \<chi>} into @{term \<chi>'}.
Moreover, it is the unique such arrow, since the components of \<open>\<phi>\<close>
are each determined by universality.
\<close>
let ?f = "A_B.MkArr (\<lambda>a. A_B.Map x' a) (\<lambda>a. A_B.Map x a) \<phi>.map"
have f_in_hom: "?f \<in> A_B.hom x' x"
proof -
have arr_f: "A_B.arr ?f"
using x' x A_B.arr_MkArr \<phi>.natural_transformation_axioms by simp
moreover have "A_B.MkIde (\<lambda>a. A_B.Map x a) = x"
using x A_B.ide_char A_B.MkArr_Map A_B.in_homE A_B.ide_in_hom by metis
moreover have "A_B.MkIde (\<lambda>a. A_B.Map x' a) = x'"
using x' A_B.ide_char A_B.MkArr_Map A_B.in_homE A_B.ide_in_hom by metis
ultimately show ?thesis
using A_B.dom_char A_B.cod_char by auto
qed
have Fun_f: "\<And>a. A.ide a \<Longrightarrow> A_B.Map ?f a = (THE fa. ?P a fa)"
using f_in_hom \<phi>.map_simp_ide by fastforce
have cones_map_f: "cones_map ?f \<chi> = \<chi>'"
using AaPa Fun_f at_ide_is_diagram assms(2) x x' cone_\<chi> cone_\<chi>' f_in_hom Fun_f
cones_map_pointwise
by presburger
show "\<guillemotleft>?f : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map ?f \<chi> = \<chi>'" using f_in_hom cones_map_f by auto
show "\<And>f'. \<guillemotleft>f' : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map f' \<chi> = \<chi>' \<Longrightarrow> f' = ?f"
proof -
fix f'
assume f': "\<guillemotleft>f' : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map f' \<chi> = \<chi>'"
have 0: "\<And>a. A.ide a \<Longrightarrow>
diagram.cones_map J B (at a D) (A_B.Map f' a) (at a \<chi>) = at a \<chi>'"
using f' cone_\<chi> cone_\<chi>' cones_map_pointwise by blast
have "f' = A_B.MkArr (A_B.Dom f') (A_B.Cod f') (A_B.Map f')"
using f' A_B.MkArr_Map by auto
also have "... = ?f"
proof (intro A_B.MkArr_eqI)
show "A_B.arr (A_B.MkArr (A_B.Dom f') (A_B.Cod f') (A_B.Map f'))"
using f' calculation by blast
show 1: "A_B.Dom f' = A_B.Map x'" using f' A_B.Map_dom by auto
show 2: "A_B.Cod f' = A_B.Map x" using f' A_B.Map_cod by auto
show "A_B.Map f' = \<phi>.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation A B (A_B.Map x') (A_B.Map x) \<phi>.map" ..
show "natural_transformation A B (A_B.Map x') (A_B.Map x) (A_B.Map f')"
using f' 1 2 A_B.arr_char [of f'] by auto
show "\<And>a. A.ide a \<Longrightarrow> A_B.Map f' a = \<phi>.map a"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret Fun_f': natural_transformation A B \<open>A_B.Dom f'\<close> \<open>A_B.Cod f'\<close>
\<open>A_B.Map f'\<close>
using f' A_B.arr_char by fast
have "A_B.Map f' a \<in> B.hom (A_B.Map x' a) (A_B.Map x a)"
using a f' Fun_f'.preserves_hom A.ide_in_hom by auto
hence "?P a (A_B.Map f' a)" using a 0 [of a] by simp
moreover have "?P a (\<phi>.map a)"
using a \<phi>.map_simp_ide Fun_f AaPa by presburger
ultimately show "A_B.Map f' a = \<phi>.map a" using a EU by blast
qed
qed
qed
finally show "f' = ?f" by auto
qed
qed
qed
qed
end
context functor_category
begin
text\<open>
A functor category \<open>[A, B]\<close> has limits of shape @{term[source=true] J}
whenever @{term B} has limits of shape @{term[source=true] J}.
\<close>
lemma has_limits_of_shape_if_target_does:
assumes "category (J :: 'j comp)"
and "B.has_limits_of_shape J"
shows "has_limits_of_shape J"
proof (unfold has_limits_of_shape_def)
have "\<And>D. diagram J comp D \<Longrightarrow> (\<exists>x \<chi>. limit_cone J comp D x \<chi>)"
proof -
fix D
assume D: "diagram J comp D"
interpret J: category J using assms(1) by auto
interpret JxA: product_category J A ..
interpret D: diagram J comp D using D by auto
interpret D: diagram_in_functor_category A B J D ..
interpret Curry: currying J A B ..
text\<open>
Given diagram @{term D} in \<open>[A, B]\<close>, choose for each object \<open>a\<close>
of \<open>A\<close> a limit cone \<open>(la, \<chi>a)\<close> for \<open>at a D\<close> in \<open>B\<close>.
\<close>
let ?l = "\<lambda>a. diagram.some_limit J B (D.at a D)"
let ?\<chi> = "\<lambda>a. diagram.some_limit_cone J B (D.at a D)"
have l\<chi>: "\<And>a. A.ide a \<Longrightarrow> diagram.limit_cone J B (D.at a D) (?l a) (?\<chi> a)"
using B.has_limits_of_shape_def D.at_ide_is_diagram assms(2)
diagram.limit_cone_some_limit_cone
by blast
text\<open>
The choice of limit cones induces a limit functor from \<open>A\<close> to \<open>B\<close>.
\<close>
interpret uncurry_D: diagram JxA.comp B "Curry.uncurry D"
proof -
interpret "functor" JxA.comp B \<open>Curry.uncurry D\<close>
using D.functor_axioms Curry.uncurry_preserves_functors by simp
interpret binary_functor J A B \<open>Curry.uncurry D\<close> ..
show "diagram JxA.comp B (Curry.uncurry D)" ..
qed
interpret uncurry_D: parametrized_diagram J A B \<open>Curry.uncurry D\<close> ..
let ?L = "uncurry_D.L ?l ?\<chi>"
let ?P = "uncurry_D.P ?l ?\<chi>"
interpret L: "functor" A B ?L
using l\<chi> uncurry_D.chosen_limits_induce_functor [of ?l ?\<chi>] by simp
have L_ide: "\<And>a. A.ide a \<Longrightarrow> ?L a = ?l a"
using uncurry_D.L_ide [of ?l ?\<chi>] l\<chi> by blast
have L_arr: "\<And>a. A.arr a \<Longrightarrow> (\<exists>!f. ?P a f) \<and> ?P a (?L a)"
using uncurry_D.L_arr [of ?l ?\<chi>] l\<chi> by blast
text\<open>
The functor \<open>L\<close> extends to a functor \<open>L'\<close> from \<open>JxA\<close>
to \<open>B\<close> that is constant on \<open>J\<close>.
\<close>
let ?L' = "\<lambda>ja. if JxA.arr ja then ?L (snd ja) else B.null"
let ?P' = "\<lambda>ja. ?P (snd ja)"
interpret L': "functor" JxA.comp B ?L'
apply unfold_locales
using L.preserves_arr L.preserves_dom L.preserves_cod
apply auto[4]
using L.preserves_comp JxA.comp_char by (elim JxA.seqE, auto)
have "\<And>ja. JxA.arr ja \<Longrightarrow> (\<exists>!f. ?P' ja f) \<and> ?P' ja (?L' ja)"
proof -
fix ja
assume ja: "JxA.arr ja"
have "A.arr (snd ja)" using ja by blast
thus "(\<exists>!f. ?P' ja f) \<and> ?P' ja (?L' ja)"
using ja L_arr by presburger
qed
hence L'_arr: "\<And>ja. JxA.arr ja \<Longrightarrow> ?P' ja (?L' ja)" by blast
have L'_ide: "\<And>ja. \<lbrakk> J.arr (fst ja); A.ide (snd ja) \<rbrakk> \<Longrightarrow> ?L' ja = ?l (snd ja)"
using L_ide l\<chi> by force
have L'_arr_map:
"\<And>ja. JxA.arr ja \<Longrightarrow> uncurry_D.P ?l ?\<chi> (snd ja) (uncurry_D.L ?l ?\<chi> (snd ja))"
using L'_arr by presburger
text\<open>
The map that takes an object \<open>(j, a)\<close> of \<open>JxA\<close> to the component
\<open>\<chi> a j\<close> of the limit cone \<open>\<chi> a\<close> is a natural transformation
from \<open>L\<close> to uncurry \<open>D\<close>.
\<close>
let ?\<chi>' = "\<lambda>ja. ?\<chi> (snd ja) (fst ja)"
interpret \<chi>': transformation_by_components JxA.comp B ?L' \<open>Curry.uncurry D\<close> ?\<chi>'
proof
fix ja
assume ja: "JxA.ide ja"
let ?j = "fst ja"
let ?a = "snd ja"
interpret \<chi>a: limit_cone J B \<open>D.at ?a D\<close> \<open>?l ?a\<close> \<open>?\<chi> ?a\<close>
using ja l\<chi> by blast
show "\<guillemotleft>?\<chi>' ja : ?L' ja \<rightarrow>\<^sub>B Curry.uncurry D ja\<guillemotright>"
using ja L'_ide [of ja] by force
next
fix ja
assume ja: "JxA.arr ja"
let ?j = "fst ja"
let ?a = "snd ja"
have j: "J.arr ?j" using ja by simp
have a: "A.arr ?a" using ja by simp
interpret D_dom_a: diagram J B \<open>D.at (A.dom ?a) D\<close>
using a D.at_ide_is_diagram by auto
interpret D_cod_a: diagram J B \<open>D.at (A.cod ?a) D\<close>
using a D.at_ide_is_diagram by auto
interpret Da: natural_transformation J B
\<open>D.at (A.dom ?a) D\<close> \<open>D.at (A.cod ?a) D\<close> \<open>D.at ?a D\<close>
using a D.functor_axioms D.functor_at_arr_is_transformation by simp
interpret \<chi>_dom_a: limit_cone J B \<open>D.at (A.dom ?a) D\<close> \<open>?l (A.dom ?a)\<close>
\<open>?\<chi> (A.dom ?a)\<close>
using a l\<chi> by simp
interpret \<chi>_cod_a: limit_cone J B \<open>D.at (A.cod ?a) D\<close> \<open>?l (A.cod ?a)\<close>
\<open>?\<chi> (A.cod ?a)\<close>
using a l\<chi> by simp
interpret Dao\<chi>_dom_a: vertical_composite J B
\<chi>_dom_a.A.map \<open>D.at (A.dom ?a) D\<close> \<open>D.at (A.cod ?a) D\<close>
\<open>?\<chi> (A.dom ?a)\<close> \<open>D.at ?a D\<close>
..
interpret Dao\<chi>_dom_a: cone J B \<open>D.at (A.cod ?a) D\<close> \<open>?l (A.dom ?a)\<close> Dao\<chi>_dom_a.map ..
show "?\<chi>' (JxA.cod ja) \<cdot>\<^sub>B ?L' ja = B (Curry.uncurry D ja) (?\<chi>' (JxA.dom ja))"
proof -
have "?\<chi>' (JxA.cod ja) \<cdot>\<^sub>B ?L' ja = ?\<chi> (A.cod ?a) (J.cod ?j) \<cdot>\<^sub>B ?L' ja"
using ja by fastforce
also have "... = D_cod_a.cones_map (?L' ja) (?\<chi> (A.cod ?a)) (J.cod ?j)"
using ja L'_arr_map [of ja] \<chi>_cod_a.cone_axioms by auto
also have "... = Dao\<chi>_dom_a.map (J.cod ?j)"
using ja \<chi>_cod_a.induced_arrowI Dao\<chi>_dom_a.cone_axioms L'_arr by presburger
also have "... = D.at ?a D (J.cod ?j) \<cdot>\<^sub>B D_dom_a.some_limit_cone (J.cod ?j)"
using ja Dao\<chi>_dom_a.map_simp_ide by fastforce
also have "... = D.at ?a D (J.cod ?j) \<cdot>\<^sub>B D.at (A.dom ?a) D ?j \<cdot>\<^sub>B ?\<chi>' (JxA.dom ja)"
using ja \<chi>_dom_a.naturality \<chi>_dom_a.ide_apex apply simp
by (metis B.comp_arr_ide \<chi>_dom_a.preserves_reflects_arr)
also have "... = (D.at ?a D (J.cod ?j) \<cdot>\<^sub>B D.at (A.dom ?a) D ?j) \<cdot>\<^sub>B ?\<chi>' (JxA.dom ja)"
using j ja B.comp_assoc by presburger
also have "... = B (D.at ?a D ?j) (?\<chi>' (JxA.dom ja))"
using a j ja Map_comp A.comp_arr_dom D.as_nat_trans.is_natural_2 by simp
also have "... = Curry.uncurry D ja \<cdot>\<^sub>B ?\<chi>' (JxA.dom ja)"
using Curry.uncurry_def by simp
finally show ?thesis by auto
qed
qed
text\<open>
Since \<open>\<chi>'\<close> is constant on \<open>J\<close>, \<open>curry \<chi>'\<close> is a cone over \<open>D\<close>.
\<close>
interpret constL: constant_functor J comp \<open>MkIde ?L\<close>
using L.as_nat_trans.natural_transformation_axioms MkArr_in_hom ide_in_hom
L.functor_axioms
by unfold_locales blast
(* TODO: This seems a little too involved. *)
have curry_L': "constL.map = Curry.curry ?L' ?L' ?L'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> constL.map j = Curry.curry ?L' ?L' ?L' j"
using Curry.curry_def constL.is_extensional by simp
moreover have "J.arr j \<Longrightarrow> constL.map j = Curry.curry ?L' ?L' ?L' j"
using Curry.curry_def constL.value_is_ide in_homE ide_in_hom by auto
ultimately show "constL.map j = Curry.curry ?L' ?L' ?L' j" by blast
qed
hence uncurry_constL: "Curry.uncurry constL.map = ?L'"
using L'.as_nat_trans.natural_transformation_axioms Curry.uncurry_curry by simp
interpret curry_\<chi>': natural_transformation J comp constL.map D
\<open>Curry.curry ?L' (Curry.uncurry D) \<chi>'.map\<close>
proof -
have "Curry.curry (Curry.uncurry D) (Curry.uncurry D) (Curry.uncurry D) = D"
using Curry.curry_uncurry D.functor_axioms D.as_nat_trans.natural_transformation_axioms
by blast
thus "natural_transformation J comp constL.map D
(Curry.curry ?L' (Curry.uncurry D) \<chi>'.map)"
using Curry.curry_preserves_transformations curry_L' \<chi>'.natural_transformation_axioms
by force
qed
interpret curry_\<chi>': cone J comp D \<open>MkIde ?L\<close> \<open>Curry.curry ?L' (Curry.uncurry D) \<chi>'.map\<close>
..
text\<open>
The value of \<open>curry_\<chi>'\<close> at each object \<open>a\<close> of \<open>A\<close> is the
limit cone \<open>\<chi> a\<close>, hence \<open>curry_\<chi>'\<close> is a limit cone.
\<close>
have 1: "\<And>a. A.ide a \<Longrightarrow> D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) = ?\<chi> a"
proof -
fix a
assume a: "A.ide a"
have "D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) =
(\<lambda>j. Curry.uncurry (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) (j, a))"
using a by simp
moreover have "... = (\<lambda>j. \<chi>'.map (j, a))"
using a Curry.uncurry_curry \<chi>'.natural_transformation_axioms by simp
moreover have "... = ?\<chi> a"
proof (intro NaturalTransformation.eqI)
interpret \<chi>a: limit_cone J B \<open>D.at a D\<close> \<open>?l a\<close> \<open>?\<chi> a\<close> using a l\<chi> by simp
interpret \<chi>': binary_functor_transformation J A B ?L' \<open>Curry.uncurry D\<close> \<chi>'.map ..
show "natural_transformation J B \<chi>a.A.map (D.at a D) (?\<chi> a)" ..
show "natural_transformation J B \<chi>a.A.map (D.at a D) (\<lambda>j. \<chi>'.map (j, a))"
proof -
have "\<chi>a.A.map = (\<lambda>j. ?L' (j, a))"
using a \<chi>a.A.map_def L'_ide by auto
thus ?thesis
using a \<chi>'.fixing_ide_gives_natural_transformation_2 by simp
qed
fix j
assume j: "J.ide j"
show "\<chi>'.map (j, a) = ?\<chi> a j"
using a j \<chi>'.map_simp_ide by simp
qed
ultimately show "D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) = ?\<chi> a" by simp
qed
hence 2: "\<And>a. A.ide a \<Longrightarrow> diagram.limit_cone J B (D.at a D) (?l a)
(D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map))"
using l\<chi> by simp
hence "limit_cone J comp D (MkIde ?L) (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map)"
using 1 2 L.functor_axioms L_ide curry_\<chi>'.cone_axioms curry_L'
D.cone_is_limit_if_pointwise_limit
by simp
thus "\<exists>x \<chi>. limit_cone J comp D x \<chi>" by blast
qed
thus "\<forall>D. diagram J comp D \<longrightarrow> (\<exists>x \<chi>. limit_cone J comp D x \<chi>)" by blast
qed
lemma has_limits_if_target_does:
assumes "B.has_limits (undefined :: 'j)"
shows "has_limits (undefined :: 'j)"
using assms B.has_limits_def has_limits_def has_limits_of_shape_if_target_does by fast
end
section "The Yoneda Functor Preserves Limits"
text\<open>
In this section, we show that the Yoneda functor from \<open>C\<close> to \<open>[Cop, S]\<close>
preserves limits.
\<close>
context yoneda_functor
begin
lemma preserves_limits:
fixes J :: "'j comp"
assumes "diagram J C D" and "diagram.has_as_limit J C D a"
shows "diagram.has_as_limit J Cop_S.comp (map o D) (map a)"
proof -
text\<open>
The basic idea of the proof is as follows:
If \<open>\<chi>\<close> is a limit cone in \<open>C\<close>, then for every object \<open>a'\<close>
of \<open>Cop\<close> the evaluation of \<open>Y o \<chi>\<close> at \<open>a'\<close> is a limit cone
in \<open>S\<close>. By the results on limits in functor categories,
this implies that \<open>Y o \<chi>\<close> is a limit cone in \<open>[Cop, S]\<close>.
\<close>
interpret J: category J using assms(1) diagram_def by auto
interpret D: diagram J C D using assms(1) by auto
from assms(2) obtain \<chi> where \<chi>: "D.limit_cone a \<chi>" by blast
interpret \<chi>: limit_cone J C D a \<chi> using \<chi> by auto
have a: "C.ide a" using \<chi>.ide_apex by auto
interpret YoD: diagram J Cop_S.comp \<open>map o D\<close>
using D.diagram_axioms functor_axioms preserves_diagrams [of J D] by simp
interpret YoD: diagram_in_functor_category Cop.comp S J \<open>map o D\<close> ..
interpret Yo\<chi>: cone J Cop_S.comp \<open>map o D\<close> \<open>map a\<close> \<open>map o \<chi>\<close>
using \<chi>.cone_axioms preserves_cones by blast
have "\<And>a'. C.ide a' \<Longrightarrow>
limit_cone J S (YoD.at a' (map o D))
(Cop_S.Map (map a) a') (YoD.at a' (map o \<chi>))"
proof -
fix a'
assume a': "C.ide a'"
interpret A': constant_functor J C a'
using a' by (unfold_locales, auto)
interpret YoD_a': diagram J S \<open>YoD.at a' (map o D)\<close>
using a' YoD.at_ide_is_diagram by simp
interpret Yo\<chi>_a': cone J S \<open>YoD.at a' (map o D)\<close>
\<open>Cop_S.Map (map a) a'\<close> \<open>YoD.at a' (map o \<chi>)\<close>
using a' YoD.cone_at_ide_is_cone Yo\<chi>.cone_axioms by fastforce
have eval_at_ide: "\<And>j. J.ide j \<Longrightarrow> YoD.at a' (map \<circ> D) j = Hom.map (a', D j)"
proof -
fix j
assume j: "J.ide j"
have "YoD.at a' (map \<circ> D) j = Cop_S.Map (map (D j)) a'"
using a' j YoD.at_simp YoD.preserves_arr [of j] by auto
also have "... = Y (D j) a'" using Y_def by simp
also have "... = Hom.map (a', D j)" using a' j D.preserves_arr by simp
finally show "YoD.at a' (map \<circ> D) j = Hom.map (a', D j)" by auto
qed
have eval_at_arr: "\<And>j. J.arr j \<Longrightarrow> YoD.at a' (map \<circ> \<chi>) j = Hom.map (a', \<chi> j)"
proof -
fix j
assume j: "J.arr j"
have "YoD.at a' (map \<circ> \<chi>) j = Cop_S.Map ((map o \<chi>) j) a'"
using a' j YoD.at_simp [of a' j "map o \<chi>"] preserves_arr by fastforce
also have "... = Y (\<chi> j) a'" using Y_def by simp
also have "... = Hom.map (a', \<chi> j)" using a' j by simp
finally show "YoD.at a' (map \<circ> \<chi>) j = Hom.map (a', \<chi> j)" by auto
qed
have Fun_map_a_a': "Cop_S.Map (map a) a' = Hom.map (a', a)"
using a a' map_simp preserves_arr [of a] by simp
show "limit_cone J S (YoD.at a' (map o D))
(Cop_S.Map (map a) a') (YoD.at a' (map o \<chi>))"
proof
fix x \<sigma>
assume \<sigma>: "YoD_a'.cone x \<sigma>"
interpret \<sigma>: cone J S \<open>YoD.at a' (map o D)\<close> x \<sigma> using \<sigma> by auto
have x: "S.ide x" using \<sigma>.ide_apex by simp
text\<open>
For each object \<open>j\<close> of \<open>J\<close>, the component \<open>\<sigma> j\<close>
is an arrow in \<open>S.hom x (Hom.map (a', D j))\<close>.
Each element \<open>e \<in> S.set x\<close> therefore determines an arrow
\<open>\<psi> (a', D j) (S.Fun (\<sigma> j) e) \<in> C.hom a' (D j)\<close>.
These arrows are the components of a cone \<open>\<kappa> e\<close> over @{term D}
with apex @{term a'}.
\<close>
have \<sigma>j: "\<And>j. J.ide j \<Longrightarrow> \<guillemotleft>\<sigma> j : x \<rightarrow>\<^sub>S Hom.map (a', D j)\<guillemotright>"
using eval_at_ide \<sigma>.preserves_hom J.ide_in_hom by force
have \<kappa>: "\<And>e. e \<in> S.set x \<Longrightarrow>
transformation_by_components
J C A'.map D (\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e))"
proof -
fix e
assume e: "e \<in> S.set x"
show "transformation_by_components J C A'.map D (\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e))"
proof
fix j
assume j: "J.ide j"
show "\<guillemotleft>\<psi> (a', D j) (S.Fun (\<sigma> j) e) : A'.map j \<rightarrow> D j\<guillemotright>"
using e j S.Fun_mapsto [of "\<sigma> j"] A'.preserves_ide Hom.set_map eval_at_ide
Hom.\<psi>_mapsto [of "A'.map j" "D j"]
by force
next
fix j
assume j: "J.arr j"
show "\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<cdot> A'.map j =
D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)"
proof -
have "\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<cdot> A'.map j =
\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<cdot> a'"
using A'.map_simp j by simp
also have "... = \<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e)"
proof -
have "\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<in> C.hom a' (D (J.cod j))"
using a' e j Hom.\<psi>_mapsto [of "A'.map j" "D (J.cod j)"] A'.map_simp
S.Fun_mapsto [of "\<sigma> (J.cod j)"] Hom.set_map eval_at_ide
by auto
thus ?thesis
using C.comp_arr_dom by fastforce
qed
also have "... = \<psi> (a', D (J.cod j)) (S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e))"
proof -
have "S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e) =
(S.Fun (Y (D j) a') o S.Fun (\<sigma> (J.dom j))) e"
by simp
also have "... = S.Fun (Y (D j) a' \<cdot>\<^sub>S \<sigma> (J.dom j)) e"
using a' e j Y_arr_ide(1) S.in_homE \<sigma>j eval_at_ide S.Fun_comp by force
also have "... = S.Fun (\<sigma> (J.cod j)) e"
using a' j x \<sigma>.is_natural_2 \<sigma>.A.map_simp S.comp_arr_dom J.arr_cod_iff_arr
J.cod_cod YoD.preserves_arr \<sigma>.is_natural_1 YoD.at_simp
by auto
finally have
"S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e) = S.Fun (\<sigma> (J.cod j)) e"
by auto
thus ?thesis by simp
qed
also have "... = D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)"
proof -
have "S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e) =
\<phi> (a', D (J.cod j)) (D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e))"
proof -
have "S.Fun (\<sigma> (J.dom j)) e \<in> Hom.set (a', D (J.dom j))"
using a' e j \<sigma>j S.Fun_mapsto [of "\<sigma> (J.dom j)"] Hom.set_map
YoD.at_simp eval_at_ide
by auto
moreover have "C.arr (\<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)) \<and>
C.dom (\<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)) = a'"
using a' e j \<sigma>j S.Fun_mapsto [of "\<sigma> (J.dom j)"] Hom.set_map eval_at_ide
Hom.\<psi>_mapsto [of a' "D (J.dom j)"]
by auto
ultimately show ?thesis
using a' e j Hom.Fun_map C.comp_arr_dom by force
qed
moreover have "D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)
\<in> C.hom a' (D (J.cod j))"
proof -
have "\<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e) \<in> C.hom a' (D (J.dom j))"
using a' e j Hom.\<psi>_mapsto [of a' "D (J.dom j)"] eval_at_ide
S.Fun_mapsto [of "\<sigma> (J.dom j)"] Hom.set_map
by auto
thus ?thesis using j D.preserves_hom by blast
qed
ultimately show ?thesis using a' j Hom.\<psi>_\<phi> by simp
qed
finally show ?thesis by auto
qed
qed
qed
let ?\<kappa> = "\<lambda>e. transformation_by_components.map J C A'.map
(\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e))"
have cone_\<kappa>e: "\<And>e. e \<in> S.set x \<Longrightarrow> D.cone a' (?\<kappa> e)"
proof -
fix e
assume e: "e \<in> S.set x"
interpret \<kappa>e: transformation_by_components J C A'.map D
\<open>\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e)\<close>
using e \<kappa> by blast
show "D.cone a' (?\<kappa> e)" ..
qed
text\<open>
Since \<open>\<kappa> e\<close> is a cone for each element \<open>e\<close> of \<open>S.set x\<close>,
by the universal property of the limit cone \<open>\<chi>\<close> there is a unique arrow
\<open>fe \<in> C.hom a' a\<close> that transforms \<open>\<chi>\<close> to \<open>\<kappa> e\<close>.
\<close>
have ex_fe: "\<And>e. e \<in> S.set x \<Longrightarrow> \<exists>!fe. \<guillemotleft>fe : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map fe \<chi> = ?\<kappa> e"
using cone_\<kappa>e \<chi>.is_universal by simp
text\<open>
The map taking \<open>e \<in> S.set x\<close> to \<open>fe \<in> C.hom a' a\<close>
determines an arrow \<open>f \<in> S.hom x (Hom (a', a))\<close> that
transforms the cone obtained by evaluating \<open>Y o \<chi>\<close> at \<open>a'\<close>
to the cone \<open>\<sigma>\<close>.
\<close>
let ?f = "S.mkArr (S.set x) (Hom.set (a', a))
(\<lambda>e. \<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)))"
have 0: "(\<lambda>e. \<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e))) \<in> S.set x \<rightarrow> Hom.set (a', a)"
proof
fix e
assume e: "e \<in> S.set x"
interpret \<kappa>e: cone J C D a' \<open>?\<kappa> e\<close> using e cone_\<kappa>e by simp
have "\<chi>.induced_arrow a' (?\<kappa> e) \<in> C.hom a' a"
using a a' e ex_fe \<chi>.induced_arrowI \<kappa>e.cone_axioms by simp
thus "\<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)) \<in> Hom.set (a', a)"
using a a' Hom.\<phi>_mapsto by auto
qed
have f: "\<guillemotleft>?f : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright>"
proof -
have "(\<lambda>e. \<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e))) \<in> S.set x \<rightarrow> Hom.set (a', a)"
proof
fix e
assume e: "e \<in> S.set x"
interpret \<kappa>e: cone J C D a' \<open>?\<kappa> e\<close> using e cone_\<kappa>e by simp
have "\<chi>.induced_arrow a' (?\<kappa> e) \<in> C.hom a' a"
using a a' e ex_fe \<chi>.induced_arrowI \<kappa>e.cone_axioms by simp
thus "\<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)) \<in> Hom.set (a', a)"
using a a' Hom.\<phi>_mapsto by auto
qed
moreover have "setp (Hom.set (a', a))"
using a a' Hom.small_homs
by (metis Fun_map_a_a' Hom.map_ide S.arr_mkIde S.ideD(1) Yo\<chi>_a'.ide_apex)
ultimately show ?thesis
using a a' x \<sigma>.ide_apex S.mkArr_in_hom [of "S.set x" "Hom.set (a', a)"]
Hom.set_subset_Univ S.mkIde_set
by simp
qed
have "YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) = \<sigma>"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J S \<sigma>.A.map (YoD.at a' (map o D)) \<sigma>"
using \<sigma>.natural_transformation_axioms by auto
have 1: "S.cod ?f = Cop_S.Map (map a) a'"
using f Fun_map_a_a' by force
interpret YoD_a'of: cone J S \<open>YoD.at a' (map o D)\<close> x
\<open>YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>))\<close>
proof -
have "YoD_a'.cone (S.cod ?f) (YoD.at a' (map o \<chi>))"
using a a' f Yo\<chi>_a'.cone_axioms preserves_arr [of a] by auto
hence "YoD_a'.cone (S.dom ?f) (YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)))"
using f YoD_a'.cones_map_mapsto S.arrI by blast
thus "cone J S (YoD.at a' (map o D)) x
(YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)))"
using f by auto
qed
show "natural_transformation J S \<sigma>.A.map (YoD.at a' (map o D))
(YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)))" ..
fix j
assume j: "J.ide j"
have "YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) j = YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f"
using f j Fun_map_a_a' Yo\<chi>_a'.cone_axioms by fastforce
also have "... = \<sigma> j"
- proof (intro S.arr_eqI)
+ proof (intro S.arr_eqI\<^sub>S\<^sub>C)
show "S.par (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) (\<sigma> j)"
using 1 f j x YoD_a'.preserves_hom by fastforce
show "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) = S.Fun (\<sigma> j)"
proof
fix e
have "e \<notin> S.set x \<Longrightarrow> S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
using 1 f j x S.Fun_mapsto [of "\<sigma> j"] \<sigma>.A.map_simp
extensional_arb [of "S.Fun (\<sigma> j)"]
by auto
moreover have "e \<in> S.set x \<Longrightarrow>
S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
proof -
assume e: "e \<in> S.set x"
interpret \<kappa>e: transformation_by_components J C A'.map D
\<open>\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e)\<close>
using e \<kappa> by blast
interpret \<kappa>e: cone J C D a' \<open>?\<kappa> e\<close> using e cone_\<kappa>e by simp
have induced_arrow: "\<chi>.induced_arrow a' (?\<kappa> e) \<in> C.hom a' a"
using a a' e ex_fe \<chi>.induced_arrowI \<kappa>e.cone_axioms by simp
have "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e =
restrict (S.Fun (YoD.at a' (map o \<chi>) j) o S.Fun ?f) (S.set x) e"
using 1 e f j S.Fun_comp YoD_a'.preserves_hom by force
also have "... = (\<phi> (a', D j) o C (\<chi> j) o \<psi> (a', a)) (S.Fun ?f e)"
using j a' f e Hom.map_simp_2 S.Fun_mkArr Hom.preserves_arr [of "(a', \<chi> j)"]
eval_at_arr
by (elim S.in_homE, auto)
also have "... = (\<phi> (a', D j) o C (\<chi> j) o \<psi> (a', a))
(\<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)))"
using e f S.Fun_mkArr by fastforce
also have "... = \<phi> (a', D j) (D.cones_map (\<chi>.induced_arrow a' (?\<kappa> e)) \<chi> j)"
using a a' e j 0 Hom.\<psi>_\<phi> induced_arrow \<chi>.cone_axioms by auto
also have "... = \<phi> (a', D j) (?\<kappa> e j)"
using \<chi>.induced_arrowI \<kappa>e.cone_axioms by fastforce
also have "... = \<phi> (a', D j) (\<psi> (a', D j) (S.Fun (\<sigma> j) e))"
using j \<kappa>e.map_def [of j] by simp
also have "... = S.Fun (\<sigma> j) e"
proof -
have "S.Fun (\<sigma> j) e \<in> Hom.set (a', D j)"
using a' e j S.Fun_mapsto [of "\<sigma> j"] eval_at_ide Hom.set_map by auto
thus ?thesis
using a' j Hom.\<phi>_\<psi> C.ide_in_hom J.ide_in_hom by blast
qed
finally show "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
by auto
qed
ultimately show "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
by auto
qed
qed
finally show "YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) j = \<sigma> j" by auto
qed
hence ff: "?f \<in> S.hom x (Hom.map (a', a)) \<and>
YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) = \<sigma>"
using f by auto
text\<open>
Any other arrow \<open>f' \<in> S.hom x (Hom.map (a', a))\<close> that
transforms the cone obtained by evaluating \<open>Y o \<chi>\<close> at @{term a'}
to the cone @{term \<sigma>}, must equal \<open>f\<close>, showing that \<open>f\<close>
is unique.
\<close>
moreover have "\<And>f'. \<guillemotleft>f' : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright> \<and>
YoD_a'.cones_map f' (YoD.at a' (map o \<chi>)) = \<sigma>
\<Longrightarrow> f' = ?f"
proof -
fix f'
assume f': "\<guillemotleft>f' : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright> \<and>
YoD_a'.cones_map f' (YoD.at a' (map o \<chi>)) = \<sigma>"
show "f' = ?f"
- proof (intro S.arr_eqI)
+ proof (intro S.arr_eqI\<^sub>S\<^sub>C)
show par: "S.par f' ?f" using f f' by (elim S.in_homE, auto)
show "S.Fun f' = S.Fun ?f"
proof
fix e
have "e \<notin> S.set x \<Longrightarrow> S.Fun f' e = S.Fun ?f e"
using f f' S.Fun_in_terms_of_comp by fastforce
moreover have "e \<in> S.set x \<Longrightarrow> S.Fun f' e = S.Fun ?f e"
proof -
assume e: "e \<in> S.set x"
have fe: "S.Fun ?f e \<in> Hom.set (a', a)"
using e f par by auto
have f'e: "S.Fun f' e \<in> Hom.set (a', a)"
using a a' e f' S.Fun_mapsto Hom.set_map by fastforce
have 1: "\<guillemotleft>\<psi> (a', a) (S.Fun f' e) : a' \<rightarrow> a\<guillemotright>"
using a a' e f' f'e S.Fun_mapsto Hom.\<psi>_mapsto Hom.set_map by blast
have 2: "\<guillemotleft>\<psi> (a', a) (S.Fun ?f e) : a' \<rightarrow> a\<guillemotright>"
using a a' e f' fe S.Fun_mapsto Hom.\<psi>_mapsto Hom.set_map by blast
interpret \<chi>ofe: cone J C D a' \<open>D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi>\<close>
proof -
have "D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<in> D.cones a \<rightarrow> D.cones a'"
using 2 D.cones_map_mapsto [of "\<psi> (a', a) (S.Fun ?f e)"]
by (elim C.in_homE, auto)
thus "cone J C D a' (D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi>)"
using \<chi>.cone_axioms by blast
qed
have A: "\<And>h j. h \<in> C.hom a' a \<Longrightarrow> J.arr j \<Longrightarrow>
S.Fun (YoD.at a' (map o \<chi>) j) (\<phi> (a', a) h)
= \<phi> (a', D (J.cod j)) (\<chi> j \<cdot> h)"
proof -
fix h j
assume j: "J.arr j"
assume h: "h \<in> C.hom a' a"
have "S.Fun (YoD.at a' (map o \<chi>) j) (\<phi> (a', a) h)
= (\<phi> (a', D (J.cod j)) \<circ> C (\<chi> j) \<circ> \<psi> (a', a)) (\<phi> (a', a) h)"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j)
= restrict (\<phi> (a', D (J.cod j)) \<circ> C (\<chi> j) \<circ> \<psi> (a', a))
(Hom.set (a', a))"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j) = S.Fun (Y (\<chi> j) a')"
using a' j YoD.at_simp Y_def Yo\<chi>.preserves_reflects_arr [of j]
by simp
also have "... = restrict (\<phi> (a', D (J.cod j)) \<circ> C (\<chi> j) \<circ> \<psi> (a', a))
(Hom.set (a', a))"
using a' j \<chi>.preserves_hom [of j "J.dom j" "J.cod j"]
Y_arr_ide [of a' "\<chi> j" a "D (J.cod j)"] \<chi>.A.map_simp S.Fun_mkArr
by fastforce
finally show ?thesis by blast
qed
thus ?thesis
using a a' h Hom.\<phi>_mapsto by auto
qed
also have "... = \<phi> (a', D (J.cod j)) (\<chi> j \<cdot> h)"
using a a' h Hom.\<psi>_\<phi> by simp
finally show "S.Fun (YoD.at a' (map o \<chi>) j) (\<phi> (a', a) h)
= \<phi> (a', D (J.cod j)) (\<chi> j \<cdot> h)"
by auto
qed
have "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
using 1 2 \<chi>.cone_axioms by (elim C.in_homE, auto)
moreover have "J.arr j \<Longrightarrow> D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
proof -
assume j: "J.arr j"
have "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
\<chi> j \<cdot> \<psi> (a', a) (S.Fun f' e)"
using j 1 \<chi>.cone_axioms by auto
also have "... = \<psi> (a', D (J.cod j)) (S.Fun (\<sigma> j) e)"
proof -
have "\<psi> (a', D (J.cod j)) (S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun f' e)) =
\<psi> (a', D (J.cod j))
(\<phi> (a', D (J.cod j)) (\<chi> j \<cdot> \<psi> (a', a) (S.Fun f' e)))"
using j a a' f'e A Hom.\<phi>_\<psi> Hom.\<psi>_mapsto by force
moreover have "\<chi> j \<cdot> \<psi> (a', a) (S.Fun f' e) \<in> C.hom a' (D (J.cod j))"
using a a' j f'e Hom.\<psi>_mapsto \<chi>.preserves_hom [of j "J.dom j" "J.cod j"]
\<chi>.A.map_simp
by auto
moreover have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun f' e) =
S.Fun (\<sigma> j) e"
using Fun_map_a_a' a a' j f' e x Yo\<chi>_a'.A.map_simp eval_at_ide
Yo\<chi>_a'.cone_axioms
by auto
ultimately show ?thesis
using a a' Hom.\<psi>_\<phi> by auto
qed
also have "... = \<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e)"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e) =
\<phi> (a', D (J.cod j)) (\<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e))"
using j a a' fe A [of "\<psi> (a', a) (S.Fun ?f e)" j] Hom.\<phi>_\<psi> Hom.\<psi>_mapsto
by auto
hence "\<psi> (a', D (J.cod j)) (S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e)) =
\<psi> (a', D (J.cod j))
(\<phi> (a', D (J.cod j)) (\<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e)))"
by simp
moreover have "\<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e) \<in> C.hom a' (D (J.cod j))"
using a a' j fe Hom.\<psi>_mapsto \<chi>.preserves_hom [of j "J.dom j" "J.cod j"]
\<chi>.A.map_simp
by auto
moreover have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e) =
S.Fun (\<sigma> j) e"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e)
= (S.Fun (YoD.at a' (map o \<chi>) j) o S.Fun ?f) e"
by simp
also have "... = S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e"
using Fun_map_a_a' a a' j f e x Yo\<chi>_a'.A.map_simp eval_at_ide
by auto
also have "... = S.Fun (\<sigma> j) e"
proof -
have "YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f =
YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) j"
using j f Yo\<chi>_a'.cone_axioms Fun_map_a_a' by auto
thus ?thesis using j ff by argo
qed
finally show ?thesis by auto
qed
ultimately show ?thesis
using a a' Hom.\<psi>_\<phi> by auto
qed
also have "... = D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
using j 2 \<chi>.cone_axioms by force
finally show "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
by auto
qed
ultimately show "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
by auto
qed
hence "\<psi> (a', a) (S.Fun f' e) = \<psi> (a', a) (S.Fun ?f e)"
using 1 2 \<chi>ofe.cone_axioms \<chi>.cone_axioms \<chi>.is_universal by blast
hence "\<phi> (a', a) (\<psi> (a', a) (S.Fun f' e)) = \<phi> (a', a) (\<psi> (a', a) (S.Fun ?f e))"
by simp
thus "S.Fun f' e = S.Fun ?f e"
using a a' fe f'e Hom.\<phi>_\<psi> by force
qed
ultimately show "S.Fun f' e = S.Fun ?f e" by auto
qed
qed
qed
ultimately have "\<exists>!f. \<guillemotleft>f : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright> \<and>
YoD_a'.cones_map f (YoD.at a' (map o \<chi>)) = \<sigma>"
using ex1I [of "\<lambda>f. S.in_hom x (Hom.map (a', a)) f \<and>
YoD_a'.cones_map f (YoD.at a' (map o \<chi>)) = \<sigma>"]
by blast
thus "\<exists>!f. \<guillemotleft>f : x \<rightarrow>\<^sub>S Cop_S.Map (map a) a'\<guillemotright> \<and>
YoD_a'.cones_map f (YoD.at a' (map o \<chi>)) = \<sigma>"
using a a' Y_def by simp
qed
qed
thus "YoD.has_as_limit (map a)"
using YoD.cone_is_limit_if_pointwise_limit Yo\<chi>.cone_axioms by auto
qed
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,1347 +1,1366 @@
(* Title: SetCat
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2016
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter SetCat
theory SetCat
imports SetCategory ConcreteCategory
begin
text\<open>
This theory proves the consistency of the \<open>set_category\<close> 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 \<open>MkArr A B F\<close>, where \<open>A\<close> and \<open>B\<close> are sets
and \<open>F\<close> is an extensional function that maps \<open>A\<close> to \<open>B\<close>.
\<close>
text\<open>
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.
An additional parameter, @{term Setp}, allows some control over which subsets of
the element type are assumed to correspond to objects of the category.
\<close>
locale setcat =
fixes dummy :: 'e
and Setp :: "'e set \<Rightarrow> bool"
assumes Setp_singleton: "Setp {x}"
and Setp_respects_subset: "A' \<subseteq> A \<Longrightarrow> Setp A \<Longrightarrow> Setp A'"
and union_preserves_Setp: "\<lbrakk> Setp A; Setp B \<rbrakk> \<Longrightarrow> Setp (A \<union> B)"
begin
lemma finite_imp_Setp: "finite A \<Longrightarrow> Setp A"
using Setp_singleton
by (metis finite_induct insert_is_Un Setp_respects_subset singleton_insert_inj_eq
union_preserves_Setp)
type_synonym 'b arr = "('b set, 'b \<Rightarrow> 'b) concrete_category.arr"
interpretation S: concrete_category \<open>Collect Setp\<close> \<open>\<lambda>A B. extensional A \<inter> (A \<rightarrow> B)\<close>
\<open>\<lambda>A. \<lambda>x \<in> A. x\<close> \<open>\<lambda>C B A g f. compose A g f\<close>
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 "\<cdot>" 55)
where "comp \<equiv> S.COMP"
notation S.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
lemma is_category:
shows "category comp"
using S.category_axioms by simp
lemma MkArr_expansion:
assumes "S.arr f"
shows "f = S.MkArr (S.Dom f) (S.Cod f) (\<lambda>x \<in> S.Dom f. S.Map f x)"
proof (intro S.arr_eqI)
show "S.arr f" by fact
show "S.arr (S.MkArr (S.Dom f) (S.Cod f) (\<lambda>x \<in> S.Dom f. S.Map f x))"
using assms S.arr_char
by (metis (mono_tags, lifting) Int_iff S.MkArr_Map extensional_restrict)
show "S.Dom f = S.Dom (S.MkArr (S.Dom f) (S.Cod f) (\<lambda>x \<in> S.Dom f. S.Map f x))"
by simp
show "S.Cod f = S.Cod (S.MkArr (S.Dom f) (S.Cod f) (\<lambda>x \<in> S.Dom f. S.Map f x))"
by simp
show "S.Map f = S.Map (S.MkArr (S.Dom f) (S.Cod f) (\<lambda>x \<in> S.Dom f. S.Map f x))"
using assms S.arr_char
by (metis (mono_tags, lifting) Int_iff S.MkArr_Map extensional_restrict)
qed
lemma arr_char:
shows "S.arr f \<longleftrightarrow> f \<noteq> S.Null \<and> Setp (S.Dom f) \<and> Setp (S.Cod f) \<and>
S.Map f \<in> extensional (S.Dom f) \<inter> (S.Dom f \<rightarrow> S.Cod f)"
using S.arr_char by auto
lemma terminal_char:
shows "S.terminal a \<longleftrightarrow> (\<exists>x. a = S.MkIde {x})"
proof
show "\<exists>x. a = S.MkIde {x} \<Longrightarrow> S.terminal a"
proof -
assume a: "\<exists>x. a = S.MkIde {x}"
from this obtain x where x: "a = S.MkIde {x}" by blast
have "S.terminal (S.MkIde {x})"
proof
show 1: "S.ide (S.MkIde {x})"
using finite_imp_Setp S.ide_MkIde by auto
show "\<And>a. S.ide a \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a \<rightarrow> S.MkIde {x}\<guillemotright>"
proof
fix a :: "'e setcat.arr"
assume a: "S.ide a"
show "\<guillemotleft>S.MkArr (S.Dom a) {x} (\<lambda>_\<in>S.Dom a. x) : a \<rightarrow> S.MkIde {x}\<guillemotright>"
proof
show 2: "S.arr (S.MkArr (S.Dom a) {x} (\<lambda>_ \<in> S.Dom a. x))"
using a 1 S.arr_MkArr [of "S.Dom a" "{x}"] S.ide_char by force
show "S.dom (S.MkArr (S.Dom a) {x} (\<lambda>_ \<in> S.Dom a. x)) = a"
using a 2 S.dom_MkArr by force
show "S.cod (S.MkArr (S.Dom a) {x} (\<lambda>_\<in>S.Dom a. x)) = S.MkIde {x}"
using 2 S.cod_MkArr by blast
qed
fix f :: "'e setcat.arr"
assume f: "\<guillemotleft>f : a \<rightarrow> S.MkIde {x}\<guillemotright>"
show "f = S.MkArr (S.Dom a) {x} (\<lambda>_ \<in> S.Dom a. x)"
proof -
have 1: "S.Dom f = S.Dom a \<and> S.Cod f = {x}"
using a f by (metis (mono_tags, lifting) S.Dom.simps(1) S.in_hom_char)
moreover have "S.Map f = (\<lambda>_ \<in> S.Dom a. x)"
proof
fix z
have "z \<notin> S.Dom a \<Longrightarrow> S.Map f z = (\<lambda>_ \<in> S.Dom a. x) z"
using f 1 MkArr_expansion
by (metis (mono_tags, lifting) S.Map.simps(1) S.in_homE restrict_apply)
moreover have "z \<in> S.Dom a \<Longrightarrow> S.Map f z = (\<lambda>_ \<in> S.Dom a. x) z"
using f 1 arr_char [of f] by fastforce
ultimately show "S.Map f z = (\<lambda>_ \<in> S.Dom a. x) z" by auto
qed
ultimately show ?thesis
using f MkArr_expansion [of f] by fastforce
qed
qed
qed
thus "S.terminal a" using x by simp
qed
show "S.terminal a \<Longrightarrow> \<exists>x. a = S.MkIde {x}"
proof -
assume a: "S.terminal a"
hence ide_a: "S.ide a" using S.terminal_def by auto
have 1: "\<exists>!x. x \<in> S.Dom a"
proof -
have "S.Dom a = {} \<Longrightarrow> \<not>S.terminal a"
proof -
assume "S.Dom a = {}"
hence 1: "a = S.MkIde {}"
using S.MkIde_Dom' \<open>S.ide a\<close> by fastforce
have "\<And>f. f \<in> S.hom (S.MkIde {undefined}) (S.MkIde ({} :: 'e set))
\<Longrightarrow> S.Map f \<in> {undefined} \<rightarrow> {}"
proof -
fix f
assume f: "f \<in> S.hom (S.MkIde {undefined}) (S.MkIde ({} :: 'e set))"
show "S.Map f \<in> {undefined} \<rightarrow> {}"
using f MkArr_expansion arr_char [of f] S.in_hom_char by auto
qed
hence "S.hom (S.MkIde {undefined}) a = {}" using 1 by auto
moreover have "S.ide (S.MkIde {undefined})"
using finite_imp_Setp
by (metis (mono_tags, lifting) finite.intros(1-2) S.ide_MkIde mem_Collect_eq)
ultimately show "\<not>S.terminal a" by blast
qed
moreover have "\<And>x x'. x \<in> S.Dom a \<and> x' \<in> S.Dom a \<and> x \<noteq> x' \<Longrightarrow> \<not>S.terminal a"
proof -
fix x x'
assume 1: "x \<in> S.Dom a \<and> x' \<in> S.Dom a \<and> x \<noteq> x'"
let ?f = "S.MkArr {undefined} (S.Dom a) (\<lambda>_ \<in> {undefined}. x)"
let ?f' = "S.MkArr {undefined} (S.Dom a) (\<lambda>_ \<in> {undefined}. x')"
have "S.par ?f ?f'"
proof -
have "\<guillemotleft>?f : S.MkIde {undefined} \<rightarrow> a\<guillemotright>"
proof
show 2: "S.arr ?f"
proof (intro S.arr_MkArr)
show "{undefined} \<in> {A. Setp A}"
by (simp add: finite_imp_Setp)
show "S.Dom a \<in> {A. Setp A}"
using ide_a S.ide_char by blast
show "(\<lambda>_ \<in> {undefined}. x) \<in> extensional {undefined} \<inter> ({undefined} \<rightarrow> S.Dom a)"
using 1 by blast
qed
show "S.dom ?f = S.MkIde {undefined}"
using 2 S.dom_MkArr by auto
show "S.cod ?f = a"
using 2 S.cod_MkArr ide_a by force
qed
moreover have "\<guillemotleft>?f' : S.MkIde {undefined} \<rightarrow> a\<guillemotright>"
proof
show 2: "S.arr ?f'"
proof (intro S.arr_MkArr)
show "{undefined} \<in> {A. Setp A}"
by (simp add: finite_imp_Setp)
show "S.Dom a \<in> {A. Setp A}"
using ide_a S.ide_char by blast
show "(\<lambda>_ \<in> {undefined}. x') \<in> extensional {undefined} \<inter> ({undefined} \<rightarrow> S.Dom a)"
using 1 by blast
qed
show "S.dom ?f' = S.MkIde {undefined}"
using 2 S.dom_MkArr by auto
show "S.cod ?f' = a"
using 2 S.cod_MkArr ide_a by force
qed
ultimately show ?thesis
by (metis (no_types, lifting) S.in_homE)
qed
moreover have "?f \<noteq> ?f'"
using 1 by (metis S.arr.inject restrict_apply' singletonI)
ultimately show "\<not>S.terminal a"
using S.cod_MkArr ide_a S.terminal_arr_unique [of ?f ?f'] by auto
qed
ultimately show ?thesis
using a by auto
qed
hence "S.Dom a = {THE x. x \<in> S.Dom a}"
using theI [of "\<lambda>x. x \<in> S.Dom a"] by auto
hence "a = S.MkIde {THE x. x \<in> S.Dom a}"
using a S.terminal_def by (metis (mono_tags, lifting) S.MkIde_Dom')
thus "\<exists>x. a = S.MkIde {x}"
by auto
qed
qed
definition IMG :: "'e setcat.arr \<Rightarrow> 'e setcat.arr"
where "IMG f = S.MkIde (S.Map f ` S.Dom f)"
interpretation S: set_category_data comp IMG
..
lemma terminal_unity:
shows "S.terminal S.unity"
using terminal_char S.unity_def someI_ex [of S.terminal]
by (metis (mono_tags, lifting))
text\<open>
The inverse maps @{term arr_of} and @{term elem_of} 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.
\<close>
definition arr_of :: "'e \<Rightarrow> 'e setcat.arr"
where "arr_of x \<equiv> S.MkIde {x}"
definition elem_of :: "'e setcat.arr \<Rightarrow> 'e"
where "elem_of t \<equiv> the_elem (S.Dom t)"
abbreviation U
where "U \<equiv> elem_of S.unity"
lemma arr_of_mapsto:
shows "arr_of \<in> UNIV \<rightarrow> S.Univ"
using terminal_char arr_of_def by fast
lemma elem_of_mapsto:
shows "elem_of \<in> Univ \<rightarrow> UNIV"
by auto
lemma elem_of_arr_of [simp]:
shows "elem_of (arr_of x) = x"
by (simp add: elem_of_def arr_of_def)
lemma arr_of_elem_of [simp]:
assumes "t \<in> S.Univ"
shows "arr_of (elem_of t) = t"
using assms terminal_char arr_of_def elem_of_def
by (metis (mono_tags, lifting) mem_Collect_eq elem_of_arr_of)
lemma inj_arr_of:
shows "inj arr_of"
by (metis elem_of_arr_of injI)
lemma bij_arr_of:
shows "bij_betw arr_of UNIV S.Univ"
proof (intro bij_betwI)
show "\<And>x :: 'e. elem_of (arr_of x) = x" by simp
show "\<And>t. t \<in> S.Univ \<Longrightarrow> arr_of (elem_of t) = t" by simp
show "arr_of \<in> UNIV \<rightarrow> S.Univ" using arr_of_mapsto by auto
show "elem_of \<in> Collect S.terminal \<rightarrow> UNIV" by auto
qed
lemma bij_elem_of:
shows "bij_betw elem_of S.Univ UNIV"
proof (intro bij_betwI)
show "\<And>t. t \<in> S.Univ \<Longrightarrow> arr_of (elem_of t) = t" by simp
show "\<And>x. x \<in> UNIV \<Longrightarrow> elem_of (arr_of x) = x" by simp
show "elem_of \<in> S.Univ \<rightarrow> UNIV" using elem_of_mapsto by auto
show "arr_of \<in> UNIV \<rightarrow> S.Univ" using arr_of_mapsto by auto
qed
lemma elem_of_img_arr_of_img [simp]:
shows "elem_of ` arr_of ` A = A"
by force
lemma arr_of_img_elem_of_img [simp]:
assumes "A \<subseteq> S.Univ"
shows "arr_of ` elem_of ` A = A"
using assms by force
lemma Dom_terminal:
assumes "S.terminal t"
shows "S.Dom t = {elem_of t}"
using assms arr_of_def
by (metis (mono_tags, lifting) S.Dom.simps(1) elem_of_def terminal_char the_elem_eq)
text\<open>
The image of a point @{term "p \<in> hom unity a"} is a terminal object, which is given
by the formula @{term "(arr_of o Fun p o elem_of) unity"}.
\<close>
lemma IMG_point:
assumes "\<guillemotleft>p : S.unity \<rightarrow> a\<guillemotright>"
shows "IMG \<in> S.hom S.unity a \<rightarrow> S.Univ"
and "IMG p = (arr_of o S.Map p o elem_of) S.unity"
proof -
show "IMG \<in> S.hom S.unity a \<rightarrow> S.Univ"
proof
fix f
assume f: "f \<in> S.hom S.unity a"
have "S.terminal (S.MkIde (S.Map f ` S.Dom S.unity))"
proof -
obtain u :: 'e where u: "S.unity = S.MkIde {u}"
using terminal_unity terminal_char
by (metis (mono_tags, lifting))
have "S.Map f ` S.Dom S.unity = {S.Map f u}"
using u by simp
thus ?thesis
using terminal_char by auto
qed
hence "S.MkIde (S.Map f ` S.Dom S.unity) \<in> S.Univ" by simp
moreover have "S.MkIde (S.Map f ` S.Dom S.unity) = IMG f"
using f IMG_def S.in_hom_char
by (metis (mono_tags, lifting) mem_Collect_eq)
ultimately show "IMG f \<in> S.Univ" by auto
qed
have "IMG p = S.MkIde (S.Map p ` S.Dom p)" using IMG_def by blast
also have "... = S.MkIde (S.Map p ` {U})"
using assms S.in_hom_char terminal_unity Dom_terminal
by (metis (mono_tags, lifting))
also have "... = (arr_of o S.Map p o elem_of) S.unity" by (simp add: arr_of_def)
finally show "IMG p = (arr_of o S.Map p o elem_of) S.unity" using assms by auto
qed
text\<open>
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.
\<close>
abbreviation MkElem :: "'e setcat.arr => 'e setcat.arr => 'e setcat.arr"
where "MkElem t a \<equiv> S.MkArr {U} (S.Dom a) (\<lambda>_ \<in> {U}. elem_of t)"
lemma MkElem_in_hom:
assumes "S.arr f" and "x \<in> S.Dom f"
shows "\<guillemotleft>MkElem (arr_of x) (S.dom f) : S.unity \<rightarrow> S.dom f\<guillemotright>"
proof -
have "(\<lambda>_ \<in> {U}. elem_of (arr_of x)) \<in> {U} \<rightarrow> S.Dom (S.dom f)"
using assms S.dom_char [of f] by simp
moreover have "S.MkIde {U} = S.unity"
using terminal_char terminal_unity
by (metis (mono_tags, lifting) elem_of_arr_of arr_of_def)
moreover have "S.MkIde (S.Dom (S.dom f)) = S.dom f"
using assms S.dom_char S.MkIde_Dom' S.ide_dom by blast
ultimately show ?thesis
using assms S.MkArr_in_hom [of "{U}" "S.Dom (S.dom f)" "\<lambda>_ \<in> {U}. elem_of (arr_of x)"]
by (metis (no_types, lifting) S.Dom.simps(1) S.Dom_in_Obj IntI S.arr_dom S.ideD(1)
restrict_extensional S.terminal_def terminal_unity)
qed
lemma MkElem_IMG:
assumes "p \<in> S.hom S.unity a"
shows "MkElem (IMG p) a = p"
proof -
have 0: "IMG p = arr_of (S.Map p U)"
using assms IMG_point(2) by auto
have 1: "S.Dom p = {U}"
using assms terminal_unity Dom_terminal
by (metis (mono_tags, lifting) S.in_hom_char mem_Collect_eq)
moreover have "S.Cod p = S.Dom a"
using assms
by (metis (mono_tags, lifting) S.in_hom_char mem_Collect_eq)
moreover have "S.Map p = (\<lambda>_ \<in> {U}. elem_of (IMG p))"
proof
fix e
show "S.Map p e = (\<lambda>_ \<in> {U}. elem_of (IMG p)) e"
proof -
have "S.Map p e = (\<lambda>x \<in> S.Dom p. S.Map p x) e"
using assms MkArr_expansion [of p]
by (metis (mono_tags, lifting) CollectD S.Map.simps(1) S.in_homE)
also have "... = (\<lambda>_ \<in> {U}. elem_of (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 S.MkArr_Map CollectD
by (metis (mono_tags, lifting) S.in_homE mem_Collect_eq)
qed
lemma inj_IMG:
assumes "S.ide a"
shows "inj_on IMG (S.hom S.unity a)"
proof (intro inj_onI)
fix x y
assume x: "x \<in> S.hom S.unity a"
assume y: "y \<in> S.hom S.unity a"
assume eq: "IMG x = IMG y"
show "x = y"
proof (intro S.arr_eqI)
show "S.arr x" using x by blast
show "S.arr y" using y by blast
show "S.Dom x = S.Dom y"
using x y S.in_hom_char by (metis (mono_tags, lifting) CollectD)
show "S.Cod x = S.Cod y"
using x y S.in_hom_char by (metis (mono_tags, lifting) CollectD)
show "S.Map x = S.Map y"
proof -
have "\<And>a. y \<in> S.hom S.unity a \<Longrightarrow> S.MkArr {U} (S.Dom a) (\<lambda>e\<in>{U}. elem_of (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 "S.ide a"
shows "S.set a = arr_of ` S.Dom a"
proof
show "S.set a \<subseteq> arr_of ` S.Dom a"
proof
fix t
assume "t \<in> S.set a"
from this obtain p where p: "\<guillemotleft>p : S.unity \<rightarrow> a\<guillemotright> \<and> t = IMG p"
using S.set_def by blast
have "t = (arr_of o S.Map p o elem_of) S.unity"
using p IMG_point(2) by blast
moreover have "(S.Map p o elem_of) S.unity \<in> S.Dom a"
using p arr_char S.in_hom_char Dom_terminal terminal_unity
by (metis (mono_tags, lifting) IntD2 Pi_split_insert_domain o_apply)
ultimately show "t \<in> arr_of ` S.Dom a" by simp
qed
show "arr_of ` S.Dom a \<subseteq> S.set a"
proof
fix t
assume "t \<in> arr_of ` S.Dom a"
from this obtain x where x: "x \<in> S.Dom a \<and> t = arr_of x" by blast
let ?p = "MkElem (arr_of x) a"
have p: "?p \<in> S.hom S.unity a"
using assms x MkElem_in_hom [of "S.dom a"] S.ideD(1-2) by force
moreover have "IMG ?p = t"
using p x elem_of_arr_of IMG_def arr_of_def
by (metis (no_types, lifting) S.Dom.simps(1) S.Map.simps(1) image_empty
image_insert image_restrict_eq)
ultimately show "t \<in> S.set a" using S.set_def by blast
qed
qed
lemma Map_via_comp:
assumes "S.arr f"
shows "S.Map f = (\<lambda>x \<in> S.Dom f. S.Map (f \<cdot> MkElem (arr_of x) (S.dom f)) U)"
proof
fix x
have "x \<notin> S.Dom f \<Longrightarrow> S.Map f x = (\<lambda>x \<in> S.Dom f. S.Map (f \<cdot> MkElem (arr_of x) (S.dom f)) U) x"
using assms arr_char [of f] IntD1 extensional_arb restrict_apply by fastforce
moreover have
"x \<in> S.Dom f \<Longrightarrow> S.Map f x = (\<lambda>x \<in> S.Dom f. S.Map (f \<cdot> MkElem (arr_of x) (S.dom f)) U) x"
proof -
assume x: "x \<in> S.Dom f"
let ?X = "MkElem (arr_of x) (S.dom f)"
have "\<guillemotleft>?X : S.unity \<rightarrow> S.dom f\<guillemotright>"
using assms x MkElem_in_hom by auto
moreover have "S.Dom ?X = {U} \<and> S.Map ?X = (\<lambda>_ \<in> {U}. x)"
using x by simp
ultimately have
"S.Map (f \<cdot> MkElem (arr_of x) (S.dom f)) = compose {U} (S.Map f) (\<lambda>_ \<in> {U}. x)"
using assms x S.Map_comp [of "MkElem (arr_of x) (S.dom f)" f]
by (metis (mono_tags, lifting) S.Cod.simps(1) S.Dom_dom S.arr_iff_in_hom S.seqE S.seqI')
thus ?thesis
using x by (simp add: compose_eq restrict_apply' singletonI)
qed
ultimately show "S.Map f x = (\<lambda>x \<in> S.Dom f. S.Map (f \<cdot> MkElem (arr_of x) (S.dom f)) U) x"
by auto
qed
lemma arr_eqI':
assumes "S.par f f'" and "\<And>t. \<guillemotleft>t : S.unity \<rightarrow> S.dom f\<guillemotright> \<Longrightarrow> f \<cdot> t = f' \<cdot> t"
shows "f = f'"
proof (intro S.arr_eqI)
show "S.arr f" using assms by simp
show "S.arr f'" using assms by simp
show "S.Dom f = S.Dom f'"
using assms by (metis (mono_tags, lifting) S.Dom_dom)
show "S.Cod f = S.Cod f'"
using assms by (metis (mono_tags, lifting) S.Cod_cod)
show "S.Map f = S.Map f'"
proof
have 1: "\<And>x. x \<in> S.Dom f \<Longrightarrow> \<guillemotleft>MkElem (arr_of x) (S.dom f) : S.unity \<rightarrow> S.dom f\<guillemotright>"
using MkElem_in_hom by (metis (mono_tags, lifting) assms(1))
fix x
show "S.Map f x = S.Map f' x"
using assms 1 \<open>S.Dom f = S.Dom f'\<close> by (simp add: Map_via_comp)
qed
qed
lemma Setp_elem_of_img:
assumes "A \<in> S.set ` Collect S.ide"
shows "Setp (elem_of ` A)"
proof -
obtain a where a: "S.ide a \<and> S.set a = A"
using assms by blast
have "elem_of ` S.set a = S.Dom a"
proof -
have "S.set a = arr_of ` S.Dom a"
using a set_char by blast
moreover have "elem_of ` arr_of ` S.Dom a = S.Dom a"
using elem_of_arr_of by force
ultimately show ?thesis by presburger
qed
thus ?thesis
using a S.ide_char by auto
qed
lemma set_MkIde_elem_of_img:
assumes "A \<subseteq> S.Univ" and "S.ide (S.MkIde (elem_of ` A))"
shows "S.set (S.MkIde (elem_of ` A)) = A"
proof -
have "S.Dom (S.MkIde (elem_of ` A)) = elem_of ` A"
by simp
moreover have "arr_of ` elem_of ` A = A"
using assms arr_of_elem_of by force
ultimately show ?thesis
using assms Setp_elem_of_img set_char S.ide_MkIde by auto
qed
(*
* We need this result, which characterizes what sets of terminals correspond
* to sets.
*)
lemma set_img_Collect_ide_iff:
shows "A \<in> S.set ` Collect S.ide \<longleftrightarrow> A \<subseteq> S.Univ \<and> Setp (elem_of ` A)"
proof
show "A \<in> S.set ` Collect S.ide \<Longrightarrow> A \<subseteq> S.Univ \<and> Setp (elem_of ` A)"
using set_char arr_of_mapsto Setp_elem_of_img by auto
assume A: "A \<subseteq> S.Univ \<and> Setp (elem_of ` A)"
have "S.ide (S.MkIde (elem_of ` A))"
using A S.ide_MkIde by blast
moreover have "S.set (S.MkIde (elem_of ` A)) = A"
using A calculation set_MkIde_elem_of_img by presburger
ultimately show "A \<in> S.set ` Collect S.ide" by blast
qed
text\<open>
The main result, which establishes the consistency of the \<open>set_category\<close> locale
and provides us with a way of obtaining ``set categories'' at arbitrary types.
\<close>
theorem is_set_category:
shows "set_category comp (\<lambda>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A))"
proof
show "\<exists>img :: 'e setcat.arr \<Rightarrow> 'e setcat.arr.
set_category_given_img comp img (\<lambda>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A))"
proof
show "set_category_given_img comp IMG (\<lambda>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A))"
proof
show "S.Univ \<noteq> {}" using terminal_char by blast
fix a :: "'e setcat.arr"
assume a: "S.ide a"
show "S.set a \<subseteq> S.Univ \<and> Setp (elem_of ` S.set a)"
using a set_img_Collect_ide_iff by auto
show "inj_on IMG (S.hom S.unity a)" using a inj_IMG terminal_unity by blast
next
fix t :: "'e setcat.arr"
assume t: "S.terminal t"
show "t \<in> IMG ` S.hom S.unity t"
proof -
have "t \<in> S.set t"
using t set_char [of t]
by (metis (mono_tags, lifting) S.Dom.simps(1) image_insert insertI1 arr_of_def
terminal_char S.terminal_def)
thus ?thesis
using t S.set_def [of t] by simp
qed
next
show "\<And>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A) \<Longrightarrow> \<exists>a. S.ide a \<and> S.set a = A"
using set_img_Collect_ide_iff by fast
next
fix a b :: "'e setcat.arr"
assume a: "S.ide a" and b: "S.ide b" and ab: "S.set a = S.set b"
show "a = b"
using a b ab set_char inj_arr_of inj_image_eq_iff S.dom_char S.in_homE S.ide_in_hom
by (metis (mono_tags, lifting))
next
fix f f' :: "'e setcat.arr"
assume par: "S.par f f'" and ff': "\<And>x. \<guillemotleft>x : S.unity \<rightarrow> S.dom f\<guillemotright> \<Longrightarrow> f \<cdot> x = f' \<cdot> x"
show "f = f'" using par ff' arr_eqI' by blast
next
fix a b :: "'e setcat.arr" and F :: "'e setcat.arr \<Rightarrow> 'e setcat.arr"
assume a: "S.ide a" and b: "S.ide b" and F: "F \<in> S.hom S.unity a \<rightarrow> S.hom S.unity b"
show "\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : S.unity \<rightarrow> S.dom f\<guillemotright> \<longrightarrow> f \<cdot> x = F x)"
proof
let ?f = "S.MkArr (S.Dom a) (S.Dom b) (\<lambda>x \<in> S.Dom a. S.Map (F (MkElem (arr_of x) a)) U)"
have 1: "\<guillemotleft>?f : a \<rightarrow> b\<guillemotright>"
proof -
have "(\<lambda>x \<in> S.Dom a. S.Map (F (MkElem (arr_of x) a)) U)
\<in> extensional (S.Dom a) \<inter> (S.Dom a \<rightarrow> S.Dom b)"
proof
show "(\<lambda>x \<in> S.Dom a. S.Map (F (MkElem (arr_of x) a)) U) \<in> extensional (S.Dom a)"
using a F by simp
show "(\<lambda>x \<in> S.Dom a. S.Map (F (MkElem (arr_of x) a)) U) \<in> S.Dom a \<rightarrow> S.Dom b"
proof
fix x
assume x: "x \<in> S.Dom a"
have "MkElem (arr_of x) a \<in> S.hom S.unity a"
using x a MkElem_in_hom [of a x] S.ide_char S.ideD(1-2) by force
hence 1: "F (MkElem (arr_of x) a) \<in> S.hom S.unity b"
using F by auto
moreover have "S.Dom (F (MkElem (arr_of x) a)) = {U}"
using 1 MkElem_IMG
by (metis (mono_tags, lifting) S.Dom.simps(1))
moreover have "S.Cod (F (MkElem (arr_of x) a)) = S.Dom b"
using 1 by (metis (mono_tags, lifting) CollectD S.in_hom_char)
ultimately have "S.Map (F (MkElem (arr_of x) a)) \<in> {U} \<rightarrow> S.Dom b"
using arr_char [of "F (MkElem (arr_of x) a)"] by blast
thus "S.Map (F (MkElem (arr_of x) a)) U \<in> S.Dom b" by blast
qed
qed
hence "\<guillemotleft>?f : S.MkIde (S.Dom a) \<rightarrow> S.MkIde (S.Dom b)\<guillemotright>"
using a b S.MkArr_in_hom S.ide_char by blast
thus ?thesis
using a b by simp
qed
moreover have "\<And>x. \<guillemotleft>x : S.unity \<rightarrow> S.dom ?f\<guillemotright> \<Longrightarrow> ?f \<cdot> x = F x"
proof -
fix x
assume x: "\<guillemotleft>x : S.unity \<rightarrow> S.dom ?f\<guillemotright>"
have 2: "x = MkElem (IMG x) a"
using a x 1 MkElem_IMG [of x a]
by (metis (mono_tags, lifting) S.in_homE mem_Collect_eq)
moreover have 5: "S.Dom x = {U} \<and> S.Cod x = S.Dom a \<and>
S.Map x = (\<lambda>_ \<in> {U}. elem_of (IMG x))"
using x 2
by (metis (no_types, lifting) S.Cod.simps(1) S.Dom.simps(1) S.Map.simps(1))
moreover have "S.Cod ?f = S.Dom b" using 1 by simp
ultimately have
3: "?f \<cdot> x =
S.MkArr {U} (S.Dom b) (compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)))"
by (metis (no_types, lifting) "1" S.Map.simps(1) S.comp_MkArr S.in_homE x)
have 4: "compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) = S.Map (F x)"
proof
fix y
have "y \<notin> {U} \<Longrightarrow>
compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) y = S.Map (F x) y"
proof -
assume y: "y \<notin> {U}"
have "compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) y = undefined"
using y compose_def extensional_arb by simp
also have "... = S.Map (F x) y"
proof -
have 5: "F x \<in> S.hom S.unity b" using x F 1 by fastforce
hence "S.Dom (F x) = {U}"
by (metis (mono_tags, lifting) "2" CollectD S.Dom.simps(1) S.in_hom_char x)
thus ?thesis
using x y F 5 arr_char [of "F x"] extensional_arb [of "S.Map (F x)" "{U}" y]
by (metis (mono_tags, lifting) CollectD Int_iff S.in_hom_char)
qed
ultimately show ?thesis by auto
qed
moreover have
"y \<in> {U} \<Longrightarrow>
compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) y = S.Map (F x) y"
proof -
assume y: "y \<in> {U}"
have "compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) y =
S.Map ?f (elem_of (IMG x))"
using y by (simp add: compose_eq restrict_apply')
also have "... = (\<lambda>x. S.Map (F (MkElem (arr_of x) a)) U) (elem_of (IMG x))"
proof -
have "elem_of (IMG x) \<in> S.Dom a"
using x y a 5 arr_char S.in_homE restrict_apply by force
thus ?thesis
using restrict_apply by simp
qed
also have "... = S.Map (F x) y"
using x y 1 2 MkElem_IMG [of x a] by simp
finally show
"compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) y = S.Map (F x) y"
by auto
qed
ultimately show
"compose {U} (S.Map ?f) (\<lambda>_ \<in> {U}. elem_of (IMG x)) y = S.Map (F x) y"
by auto
qed
show "?f \<cdot> x = F x"
proof (intro S.arr_eqI)
have 5: "?f \<cdot> x \<in> S.hom S.unity b" using 1 x by blast
have 6: "F x \<in> S.hom S.unity b"
using x F 1
by (metis (mono_tags, lifting) PiE S.in_homE mem_Collect_eq)
show "S.arr (comp ?f x)" using 5 by blast
show "S.arr (F x)" using 6 by blast
show "S.Dom (comp ?f x) = S.Dom (F x)"
using 5 6 by (metis (mono_tags, lifting) CollectD S.in_hom_char)
show "S.Cod (comp ?f x) = S.Cod (F x)"
using 5 6 by (metis (mono_tags, lifting) CollectD S.in_hom_char)
show "S.Map (comp ?f x) = S.Map (F x)"
using 3 4 by simp
qed
qed
thus "\<guillemotleft>?f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : S.unity \<rightarrow> S.dom ?f\<guillemotright> \<longrightarrow> comp ?f x = F x)"
using 1 by blast
qed
next
show "\<And>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A) \<Longrightarrow> A \<subseteq> S.Univ"
by simp
show "\<And>A' A. \<lbrakk>A' \<subseteq> A; A \<subseteq> S.Univ \<and> Setp (elem_of ` A)\<rbrakk> \<Longrightarrow> A' \<subseteq> S.Univ \<and> Setp (elem_of ` A')"
by (meson image_mono setcat.Setp_respects_subset setcat_axioms subset_trans)
show "\<And>A B. \<lbrakk>A \<subseteq> S.Univ \<and> Setp (elem_of ` A); B \<subseteq> S.Univ \<and> Setp (elem_of ` B)\<rbrakk>
\<Longrightarrow> A \<union> B \<subseteq> S.Univ \<and> Setp (elem_of ` (A \<union> B))"
by (simp add: image_Un union_preserves_Setp)
qed
qed
qed
text\<open>
\<open>SetCat\<close> can be viewed as a concrete set category over its own element type
@{typ 'a}, using @{term arr_of} as the required injection from @{typ 'a} to the universe
of \<open>SetCat\<close>.
\<close>
corollary is_concrete_set_category:
shows "concrete_set_category comp (\<lambda>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A)) UNIV arr_of"
proof -
interpret S': set_category comp \<open>\<lambda>A. A \<subseteq> S.Univ \<and> Setp (elem_of ` A)\<close>
using is_set_category by auto
show ?thesis
proof
show 1: "arr_of \<in> UNIV \<rightarrow> S'.Univ"
using arr_of_def terminal_char by force
show "inj_on arr_of UNIV"
using inj_arr_of by blast
qed
qed
text\<open>
As a consequence of the categoricity of the \<open>set_category\<close> axioms,
if @{term S} interprets \<open>set_category\<close>, and if @{term \<phi>} is a bijection between
the universe of @{term S} and the elements of type @{typ 'a}, then @{term S} is isomorphic
to the category \<open>setcat\<close> of @{typ 'a} sets and functions between them constructed here.
\<close>
corollary set_category_iso_SetCat:
fixes S :: "'s comp" and \<phi> :: "'s \<Rightarrow> 'e"
assumes "set_category S \<S>"
and "bij_betw \<phi> (set_category_data.Univ S) UNIV"
and "\<And>A. \<S> A \<longleftrightarrow> A \<subseteq> set_category_data.Univ S \<and> (arr_of \<circ> \<phi>) ` A \<in> S.set ` Collect S.ide"
shows "\<exists>\<Phi>. invertible_functor S comp \<Phi>
\<and> (\<forall>m. set_category.incl S \<S> m
\<longrightarrow> set_category.incl comp (\<lambda>A. A \<in> S.set ` Collect S.ide) (\<Phi> m))"
proof -
interpret S': set_category S \<S> using assms by auto
let ?\<psi> = "inv_into S'.Univ \<phi>"
have "bij_betw (arr_of o \<phi>) S'.Univ S.Univ"
proof (intro bij_betwI)
show "arr_of o \<phi> \<in> S'.Univ \<rightarrow> Collect S.terminal"
using assms(2) arr_of_mapsto by auto
show "?\<psi> o elem_of \<in> S.Univ \<rightarrow> S'.Univ"
proof
fix x :: "'e setcat.arr"
assume x: "x \<in> S.Univ"
show "(inv_into S'.Univ \<phi> \<circ> elem_of) x \<in> S'.Univ"
using x assms(2) bij_betw_def comp_apply inv_into_into
by (metis UNIV_I)
qed
fix t
assume "t \<in> S'.Univ"
thus "(?\<psi> o elem_of) ((arr_of o \<phi>) t) = t"
using assms(2) bij_betw_inv_into_left
by (metis comp_apply elem_of_arr_of)
next
fix t' :: "'e setcat.arr"
assume "t' \<in> S.Univ"
thus "(arr_of o \<phi>) ((?\<psi> o elem_of) t') = t'"
using assms(2) by (simp add: bij_betw_def f_inv_into_f)
qed
thus ?thesis
using assms is_set_category set_img_Collect_ide_iff
set_category_is_categorical
[of S \<S> comp "\<lambda>A. A \<in> S.set ` Collect S.ide" "arr_of o \<phi>"]
by simp
qed
sublocale category comp
using is_category by blast
sublocale set_category comp \<open>\<lambda>A. A \<subseteq> Collect S.terminal \<and> Setp (elem_of ` A)\<close>
using is_set_category set_img_Collect_ide_iff by simp
- sublocale concrete_set_category comp \<open>\<lambda>A. A \<subseteq> Collect S.terminal \<and> Setp (elem_of ` A)\<close>
+ interpretation concrete_set_category comp \<open>\<lambda>A. A \<subseteq> Collect S.terminal \<and> Setp (elem_of ` A)\<close>
UNIV arr_of
using is_concrete_set_category by simp
-
+
end
-
+(*
+ interpretation SetCat: setcat undefined \<open>\<lambda>S. True\<close>
+ by unfold_locales auto
+ *)
text\<open>
Here we discard the temporary interpretations \<open>S\<close>, leaving only the exported
definitions and facts.
\<close>
context setcat
begin
text\<open>
We establish mappings to pass back and forth between objects and arrows of the category
and sets and functions on the underlying elements.
\<close>
+ interpretation set_category comp \<open>\<lambda>A. A \<subseteq> Collect terminal \<and> Setp (elem_of ` A)\<close>
+ using is_set_category by blast
+ interpretation concrete_set_category comp \<open>\<lambda>A. A \<subseteq> Univ \<and> Setp (elem_of ` A)\<close> UNIV arr_of
+ using is_concrete_set_category by blast
+
definition set_of_ide :: "'e setcat.arr \<Rightarrow> 'e set"
where "set_of_ide a \<equiv> elem_of ` set a"
definition ide_of_set :: "'e set \<Rightarrow> 'e setcat.arr"
where "ide_of_set A \<equiv> mkIde (arr_of ` A)"
lemma bij_betw_ide_set:
shows "set_of_ide \<in> Collect ide \<rightarrow> Collect Setp"
and "ide_of_set \<in> Collect Setp \<rightarrow> Collect ide"
and [simp]: "ide a \<Longrightarrow> ide_of_set (set_of_ide a) = a"
and [simp]: "Setp A \<Longrightarrow> set_of_ide (ide_of_set A) = A"
and "bij_betw set_of_ide (Collect ide) (Collect Setp)"
and "bij_betw ide_of_set (Collect Setp) (Collect ide)"
proof -
show 1: "set_of_ide \<in> Collect ide \<rightarrow> Collect Setp"
using setp_set_ide set_of_ide_def by auto
show 2: "ide_of_set \<in> Collect Setp \<rightarrow> Collect ide"
proof
fix A :: "'e set"
assume A: "A \<in> Collect Setp"
have "arr_of ` A \<subseteq> Univ"
using A arr_of_mapsto by auto
moreover have "Setp (elem_of ` arr_of ` A)"
using A elem_of_arr_of Setp_respects_subset by simp
ultimately have "ide (mkIde (arr_of ` A))"
using ide_mkIde by blast
thus "ide_of_set A \<in> Collect ide"
using ide_of_set_def by simp
qed
show 3: "\<And>a. ide a \<Longrightarrow> ide_of_set (set_of_ide a) = a"
using arr_of_img_elem_of_img ide_of_set_def mkIde_set set_of_ide_def setp_set_ide
by presburger
show 4: "\<And>A. Setp A \<Longrightarrow> set_of_ide (ide_of_set A) = A"
proof -
fix A :: "'e set"
assume A: "Setp A"
have "elem_of ` set (mkIde (arr_of ` A)) = elem_of ` arr_of ` A"
proof -
have "arr_of ` A \<subseteq> Univ"
using A arr_of_mapsto by blast
moreover have "Setp (elem_of ` arr_of ` A)"
using A by simp
ultimately have "set (mkIde (arr_of ` A)) = arr_of ` A"
using A set_mkIde by blast
thus ?thesis by auto
qed
also have "... = A"
using A elem_of_arr_of by force
finally show "set_of_ide (ide_of_set A) = A"
using ide_of_set_def set_of_ide_def by simp
qed
show "bij_betw set_of_ide (Collect ide) (Collect Setp)"
using 1 2 3 4
by (intro bij_betwI) blast+
show "bij_betw ide_of_set (Collect Setp) (Collect ide)"
using 1 2 3 4
by (intro bij_betwI) blast+
qed
definition fun_of_arr :: "'e setcat.arr \<Rightarrow> 'e \<Rightarrow> 'e"
where "fun_of_arr f \<equiv> restrict (elem_of o Fun f o arr_of) (elem_of `Dom f)"
definition arr_of_fun :: "'e set \<Rightarrow> 'e set \<Rightarrow> ('e \<Rightarrow> 'e) \<Rightarrow> 'e setcat.arr"
where "arr_of_fun A B F \<equiv> mkArr (arr_of ` A) (arr_of ` B) (arr_of o F o elem_of)"
lemma bij_betw_hom_fun:
shows "fun_of_arr \<in> hom a b \<rightarrow> extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b)"
and "\<lbrakk>Setp A; Setp B\<rbrakk> \<Longrightarrow> arr_of_fun A B \<in> (A \<rightarrow> B) \<rightarrow> hom (ide_of_set A) (ide_of_set B)"
and "f \<in> hom a b \<Longrightarrow> arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f"
and "\<lbrakk>Setp A; Setp B; F \<in> A \<rightarrow> B; F \<in> extensional A\<rbrakk> \<Longrightarrow> fun_of_arr (arr_of_fun A B F) = F"
and "\<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> bij_betw fun_of_arr (hom a b)
(extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b))"
and "\<lbrakk>Setp A; Setp B\<rbrakk> \<Longrightarrow>
bij_betw (arr_of_fun A B)
(extensional A \<inter> (A \<rightarrow> B)) (hom (ide_of_set A) (ide_of_set B))"
proof -
show 1: "\<And>a b. fun_of_arr \<in>
hom a b \<rightarrow> extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b)"
proof
fix a b f
assume f: "f \<in> hom a b"
have Dom: "Dom f = arr_of ` set_of_ide a"
using f set_of_ide_def
by (metis (mono_tags, lifting) arr_dom arr_mkIde bij_betw_ide_set(3)
ide_dom ide_of_set_def in_homE mem_Collect_eq set_mkIde)
have Cod: "Cod f = arr_of ` set_of_ide b"
using f set_of_ide_def
by (metis (mono_tags, lifting) arr_cod arr_mkIde bij_betw_ide_set(3)
ide_cod ide_of_set_def in_homE mem_Collect_eq set_mkIde)
have "fun_of_arr f \<in> set_of_ide a \<rightarrow> set_of_ide b"
proof
fix x
assume x: "x \<in> set_of_ide a"
have 1: "arr_of x \<in> Dom f"
using x Dom by blast
hence "Fun f (arr_of x) \<in> Cod f"
using f Fun_mapsto Cod by blast
hence "elem_of (Fun f (arr_of x)) \<in> set_of_ide b"
using Cod by auto
hence "restrict (elem_of o Fun f o arr_of) (elem_of ` Dom f) x \<in> set_of_ide b"
using 1 by force
thus "fun_of_arr f x \<in> set_of_ide b"
unfolding fun_of_arr_def by auto
qed
moreover have "fun_of_arr f \<in> extensional (set_of_ide a)"
unfolding fun_of_arr_def
using set_of_ide_def f by blast
ultimately show "fun_of_arr f \<in> extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b)"
by blast
qed
show 2: "\<And>A B. \<lbrakk>Setp A; Setp B\<rbrakk> \<Longrightarrow>
arr_of_fun A B \<in> (A \<rightarrow> B) \<rightarrow> hom (ide_of_set A) (ide_of_set B)"
proof
fix x and A B :: "'e set"
assume A: "Setp A" and B: "Setp B"
assume x: "x \<in> A \<rightarrow> B"
show "arr_of_fun A B x \<in> hom (ide_of_set A) (ide_of_set B)"
proof
show "\<guillemotleft>arr_of_fun A B x : ide_of_set A \<rightarrow> ide_of_set B\<guillemotright>"
proof
show 1: "arr (arr_of_fun A B x)"
proof -
have "arr_of ` A \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` A)"
using A arr_of_mapsto elem_of_arr_of
by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2)
ide_implies_incl ide_of_set_def incl_def mem_Collect_eq)
moreover have "arr_of ` B \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` B)"
using B arr_of_mapsto elem_of_arr_of
by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2)
ide_implies_incl ide_of_set_def incl_def mem_Collect_eq)
moreover have "arr_of \<circ> x \<circ> elem_of \<in> arr_of ` A \<rightarrow> arr_of ` B"
using x by auto
ultimately show ?thesis
unfolding arr_of_fun_def by blast
qed
show "dom (arr_of_fun A B x) = ide_of_set A"
using 1 dom_mkArr ide_of_set_def arr_of_fun_def by simp
show "cod (arr_of_fun A B x) = ide_of_set B"
using 1 cod_mkArr ide_of_set_def arr_of_fun_def by simp
qed
qed
qed
show 3: "\<And>a b f. f \<in> hom a b \<Longrightarrow> arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f"
proof -
fix a b f
assume f: "f \<in> hom a b"
have 1: "Dom f = set a \<and> Cod f = set b"
using f by blast
have Dom: "Dom f \<subseteq> Univ \<and> Setp (elem_of ` Dom f)"
using setp_set_ide f ide_dom by blast
have Cod: "Cod f \<subseteq> Univ \<and> Setp (elem_of ` Cod f)"
using setp_set_ide f ide_cod by blast
have "mkArr (set a) (set b)
(arr_of \<circ> restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ> elem_of) =
mkArr (Dom f) (Cod f)
(arr_of \<circ> restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ> elem_of)"
using 1 by simp
also have "... = mkArr (Dom f) (Cod f) (Fun f)"
proof (intro mkArr_eqI')
show 2: "\<And>x. x \<in> Dom f \<Longrightarrow>
(arr_of \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ>
elem_of) x =
Fun f x"
proof -
fix x
assume x: "x \<in> Dom f"
hence 1: "elem_of x \<in> elem_of ` Dom f"
by blast
have "(arr_of \<circ> restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ> elem_of) x =
arr_of (restrict (elem_of o Fun f o arr_of) (elem_of ` Dom f) (elem_of x))"
by auto
also have "... = arr_of ((elem_of o Fun f o arr_of) (elem_of x))"
using 1 by auto
also have "... = arr_of (elem_of (Fun f (arr_of (elem_of x))))"
by auto
also have "... = arr_of (elem_of (Fun f x))"
using x arr_of_elem_of \<open>Dom f \<subseteq> Univ \<and> Setp (elem_of ` Dom f)\<close> by auto
also have "... = Fun f x"
using x f Dom Cod Fun_mapsto arr_of_elem_of [of "Fun f x"] by blast
finally show "(arr_of \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ>
elem_of) x =
Fun f x"
by blast
qed
have "arr_of \<circ> restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ> elem_of
\<in> Dom f \<rightarrow> Cod f"
proof
fix x
assume x: "x \<in> Dom f"
have "(arr_of \<circ> restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ> elem_of) x =
Fun f x"
using 2 x by blast
moreover have "... \<in> Cod f"
using f x Fun_mapsto by blast
ultimately show "(arr_of \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ>
elem_of) x
\<in> Cod f"
by argo
qed
thus "arr (mkArr (Dom f) (Cod f)
(arr_of \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ>
elem_of))"
using Dom Cod by blast
qed
finally have "mkArr (set a) (set b)
(arr_of \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) \<circ>
elem_of) = f"
using f mkArr_Fun
by (metis (no_types, lifting) in_homE mem_Collect_eq)
thus "arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f"
using 1 f
by (metis (no_types, lifting) Cod Dom arr_of_img_elem_of_img arr_of_fun_def
fun_of_arr_def set_of_ide_def)
qed
show 4: "\<And>A B F. \<lbrakk>Setp A; Setp B; F \<in> A \<rightarrow> B; F \<in> extensional A\<rbrakk>
\<Longrightarrow> fun_of_arr (arr_of_fun A B F) = F"
proof
fix F and A B :: "'e set"
assume A: "Setp A" and B: "Setp B"
assume F: "F \<in> A \<rightarrow> B" and ext: "F \<in> extensional A"
have 4: "arr (mkArr (arr_of ` A) (arr_of ` B) (arr_of \<circ> F \<circ> elem_of))"
proof -
have "arr_of ` A \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` A)"
using A
by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) ide_implies_incl
ide_of_set_def incl_def mem_Collect_eq)
moreover have "arr_of ` B \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` B)"
using B
by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) ide_implies_incl
ide_of_set_def incl_def mem_Collect_eq)
moreover have "arr_of \<circ> F \<circ> elem_of \<in> arr_of ` A \<rightarrow> arr_of ` B"
using F by auto
ultimately show ?thesis by blast
qed
show "\<And>x. fun_of_arr (arr_of_fun A B F) x = F x"
proof -
fix x
have "fun_of_arr (arr_of_fun A B F) x =
restrict (elem_of \<circ>
Fun (mkArr (arr_of ` A) (arr_of ` B) (arr_of \<circ> F \<circ> elem_of)) \<circ>
arr_of) A x"
proof -
have "elem_of ` Dom (mkArr (arr_of ` A) (arr_of ` B) (arr_of \<circ> F \<circ> elem_of)) = A"
using A 4 elem_of_arr_of dom_mkArr set_of_ide_def bij_betw_ide_set(4) ide_of_set_def
by auto
thus ?thesis
using arr_of_fun_def fun_of_arr_def by auto
qed
also have "... = F x"
proof (cases "x \<in> A")
show "x \<notin> A \<Longrightarrow> ?thesis"
using ext extensional_arb by fastforce
assume x: "x \<in> A"
show "restrict
(elem_of \<circ>
Fun (mkArr (arr_of ` A) (arr_of ` B) (arr_of \<circ> F \<circ> elem_of)) \<circ>
arr_of) A x =
F x"
using x 4 Fun_mkArr by simp
qed
finally show "fun_of_arr (arr_of_fun A B F) x = F x"
by blast
qed
qed
show "\<lbrakk>Setp A; Setp B\<rbrakk> \<Longrightarrow>
bij_betw (arr_of_fun A B) (extensional A \<inter> (A \<rightarrow> B))
(hom (ide_of_set A) (ide_of_set B))"
proof -
assume A: "Setp A" and B: "Setp B"
have "ide (ide_of_set A) \<and> ide (ide_of_set B)"
using A B bij_betw_ide_set(2) by auto
have "set_of_ide (ide_of_set A) = A \<and> set_of_ide (ide_of_set B) = B"
using A B by simp
show ?thesis
using A B 1 [of "ide_of_set A" "ide_of_set B"] 2 3 4
apply (intro bij_betwI)
apply auto
apply blast
by fastforce
qed
show "\<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> bij_betw fun_of_arr (hom a b)
(extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b))"
proof (intro bij_betwI)
assume a: "ide a" and b: "ide b"
show "fun_of_arr \<in> hom a b \<rightarrow> extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b)"
using 1 by blast
show "arr_of_fun (set_of_ide a) (set_of_ide b) \<in>
extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b) \<rightarrow> hom a b"
using a b 2 [of "set_of_ide a" "set_of_ide b"] setp_set_ide set_of_ide_def
bij_betw_ide_set(3)
by auto
show "\<And>f. f \<in> hom a b \<Longrightarrow> arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f"
using a b 3 by blast
show "\<And>F. F \<in> extensional (set_of_ide a) \<inter> (set_of_ide a \<rightarrow> set_of_ide b) \<Longrightarrow>
fun_of_arr (arr_of_fun (set_of_ide a) (set_of_ide b) F) = F"
using a b 4 [of "set_of_ide a" "set_of_ide b"]
by (metis (no_types, lifting) IntE set_of_ide_def setp_set_ide)
qed
qed
lemma fun_of_arr_ide:
assumes "ide a"
shows "fun_of_arr a = restrict id (elem_of ` Dom a)"
proof
fix x
show "fun_of_arr a x = restrict id (elem_of ` Dom a) x"
proof (cases "x \<in> elem_of ` Dom a")
show "x \<notin> elem_of ` Dom a \<Longrightarrow> ?thesis"
using fun_of_arr_def extensional_arb by auto
assume x: "x \<in> elem_of ` Dom a"
have "fun_of_arr a x = restrict (elem_of \<circ> Fun a \<circ> arr_of) (elem_of ` Dom a) x"
using x fun_of_arr_def by simp
also have "... = elem_of (Fun a (arr_of x))"
using x by auto
also have "... = elem_of ((\<lambda>x\<in>set a. x) (arr_of x))"
using assms x Fun_ide by auto
also have "... = elem_of (arr_of x)"
proof -
have "x \<in> elem_of ` set a"
using assms x ideD(2) by force
hence "arr_of x \<in> set a"
by (metis (mono_tags, lifting) arr_of_img_elem_of_img assms image_eqI setp_set_ide)
thus ?thesis by simp
qed
also have "... = x"
by simp
also have "... = restrict id (elem_of ` Dom a) x"
using x by auto
finally show ?thesis by blast
qed
qed
lemma arr_of_fun_id:
assumes "Setp A"
shows "arr_of_fun A A (restrict id A) = ide_of_set A"
proof -
have A: "arr_of ` A \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` A)"
using assms elem_of_arr_of
by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) ide_implies_incl
ide_of_set_def incl_def mem_Collect_eq)
have "arr_of_fun A A (restrict id A) =
mkArr (arr_of ` A) (arr_of ` A) (arr_of \<circ> restrict id A \<circ> elem_of)"
unfolding arr_of_fun_def by simp
also have "... = mkArr (arr_of ` A) (arr_of ` A) (\<lambda>x. x)"
using A arr_mkArr
by (intro mkArr_eqI') auto
also have "... = ide_of_set A"
using A ide_of_set_def mkIde_as_mkArr by simp
finally show ?thesis by blast
qed
lemma fun_of_arr_comp:
assumes "f \<in> hom a b" and "g \<in> hom b c"
shows "fun_of_arr (comp g f) = restrict (fun_of_arr g \<circ> fun_of_arr f) (set_of_ide a)"
proof -
have 1: "seq g f"
using assms by blast
have "fun_of_arr (comp g f) =
restrict (elem_of \<circ> Fun (comp g f) \<circ> arr_of) (elem_of ` Dom (comp g f))"
unfolding fun_of_arr_def by blast
also have "... = restrict (elem_of \<circ> Fun (comp g f) \<circ> arr_of) (elem_of ` Dom f)"
using assms dom_comp seqI' by auto
also have "... = restrict (elem_of \<circ> restrict (Fun g \<circ> Fun f) (Dom f) \<circ> arr_of)
(elem_of ` Dom f)"
using 1 Fun_comp by auto
also have "... = restrict (restrict (elem_of o Fun g o arr_of) (elem_of ` Dom g) \<circ>
restrict (elem_of o Fun f o arr_of) (elem_of ` Dom f))
(elem_of ` Dom f)"
proof
fix x
show "restrict (elem_of \<circ> restrict (Fun g \<circ> Fun f) (Dom f) \<circ> arr_of) (elem_of ` Dom f) x =
restrict (restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g) \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f))
(elem_of ` Dom f) x"
proof (cases "x \<in> elem_of ` Dom f")
show "x \<notin> elem_of ` Dom f \<Longrightarrow> ?thesis"
by auto
assume x: "x \<in> elem_of ` Dom f"
have 2: "arr_of x \<in> Dom f"
proof -
have "arr_of x \<in> arr_of ` elem_of ` Dom f"
using x by simp
thus ?thesis
by (metis (no_types, lifting) 1 arr_of_img_elem_of_img ide_dom seqE setp_set_ide)
qed
have 3: "Dom g = Cod f"
using assms by fastforce
have "restrict (elem_of \<circ> restrict (Fun g \<circ> Fun f) (Dom f) \<circ> arr_of) (elem_of ` Dom f) x =
elem_of (Fun g (Fun f (arr_of x)))"
using x 2 by simp
also have "... = restrict
(restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g) \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f))
(elem_of ` Dom f) x"
proof -
have "restrict (restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g) \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f))
(elem_of ` Dom f) x =
elem_of (Fun g (Fun f (arr_of x)))"
proof -
have "restrict (restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g) \<circ>
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f))
(elem_of ` Dom f) x =
(restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g) o
restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f)) x"
using 2 by force
also have "... = restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g)
(restrict (elem_of \<circ> Fun f \<circ> arr_of) (elem_of ` Dom f) x)"
by simp
also have "... = restrict (elem_of \<circ> Fun g \<circ> arr_of) (elem_of ` Dom g)
(elem_of (Fun f (arr_of x)))"
using 2 by force
also have "... = (elem_of o Fun g o arr_of) (elem_of (Fun f (arr_of x)))"
proof -
have "elem_of (Fun f (arr_of x)) \<in> elem_of ` Dom g"
using assms 2 3 Fun_mapsto [of f] by blast
thus ?thesis by simp
qed
also have "... = elem_of (Fun g (arr_of (elem_of (Fun f (arr_of x)))))"
by simp
also have "... = elem_of (Fun g (Fun f (arr_of x)))"
proof -
have "Fun f (arr_of x) \<in> Univ"
using assms 2 setp_set_ide ide_cod Fun_mapsto by blast
thus ?thesis
using 2 by simp
qed
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
qed
also have "... = restrict (fun_of_arr g o fun_of_arr f) (elem_of ` Dom f)"
unfolding fun_of_arr_def by blast
finally show ?thesis
unfolding set_of_ide_def
using assms by blast
qed
lemma arr_of_fun_comp:
assumes "Setp A" and "Setp B" and "Setp C"
and "F \<in> extensional A \<inter> (A \<rightarrow> B)" and "G \<in> extensional B \<inter> (B \<rightarrow> C)"
shows "arr_of_fun A C (G o F) = comp (arr_of_fun B C G) (arr_of_fun A B F)"
proof -
have A: "arr_of ` A \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` A)"
using assms elem_of_arr_of
by (metis (no_types, lifting) Pi_iff arr_mkIde bij_betw_ide_set(2)
ide_implies_incl ide_of_set_def incl_def mem_Collect_eq)
have B: "arr_of ` B \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` B)"
using assms elem_of_arr_of
by (metis (no_types, lifting) Pi_iff arr_mkIde bij_betw_ide_set(2)
ide_implies_incl ide_of_set_def incl_def mem_Collect_eq)
have C: "arr_of ` C \<subseteq> Univ \<and> Setp (elem_of ` arr_of ` C)"
using assms elem_of_arr_of
by (metis (no_types, lifting) Pi_iff arr_mkIde bij_betw_ide_set(2)
ide_implies_incl ide_of_set_def incl_def mem_Collect_eq)
have "arr_of_fun A C (G o F) = mkArr (arr_of ` A) (arr_of ` C) (arr_of \<circ> (G \<circ> F) \<circ> elem_of)"
unfolding arr_of_fun_def by simp
also have "... = mkArr (arr_of ` A) (arr_of ` C)
((arr_of \<circ> G \<circ> elem_of) o (arr_of o F \<circ> elem_of))"
proof (intro mkArr_eqI')
show "arr (mkArr (arr_of ` A) (arr_of ` C) (arr_of \<circ> (G \<circ> F) \<circ> elem_of))"
proof -
have "arr_of \<circ> (G \<circ> F) \<circ> elem_of \<in> arr_of ` A \<rightarrow> arr_of ` C"
using assms by force
thus ?thesis
using A B C by blast
qed
show "\<And>x. x \<in> arr_of ` A \<Longrightarrow>
(arr_of \<circ> (G \<circ> F) \<circ> elem_of) x =
((arr_of \<circ> G \<circ> elem_of) o (arr_of o F \<circ> elem_of)) x"
by simp
qed
also have "... = comp (mkArr (arr_of ` B) (arr_of ` C) (arr_of \<circ> G \<circ> elem_of))
(mkArr (arr_of ` A) (arr_of ` B) (arr_of o F o elem_of))"
proof -
have "arr (mkArr (arr_of ` B) (arr_of ` C) (arr_of \<circ> G \<circ> elem_of))"
using assms B C elem_of_arr_of by fastforce
moreover have "arr (mkArr (arr_of ` A) (arr_of ` B) (arr_of o F o elem_of))"
using assms A B elem_of_arr_of by fastforce
ultimately show ?thesis
using comp_mkArr by auto
qed
also have "... = comp (arr_of_fun B C G) (arr_of_fun A B F)"
using arr_of_fun_def by presburger
finally show ?thesis by blast
qed
end
text\<open>
When there is no restriction on the sets that determine objects, the resulting set category
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.
\<close>
locale replete_setcat =
fixes dummy :: 'e
begin
interpretation SC: setcat dummy \<open>\<lambda>_. True\<close>
- proof
- show True by blast
- qed
+ by unfold_locales blast
definition comp
where "comp \<equiv> SC.comp"
definition arr_of
where "arr_of \<equiv> SC.arr_of"
definition elem_of
where "elem_of \<equiv> SC.elem_of"
sublocale replete_set_category comp
unfolding comp_def
- using SC.set_img_Collect_ide_iff SC.set_category_given_img_axioms
- by unfold_locales auto
+ using SC.is_set_category replete_set_category_def set_category_def by auto
+
+ lemma is_replete_set_category:
+ shows "replete_set_category comp"
+ ..
+
+ lemma is_set_category\<^sub>R\<^sub>S\<^sub>C:
+ shows "set_category comp (\<lambda>A. A \<subseteq> Univ)"
+ using is_set_category by blast
sublocale concrete_set_category comp setp UNIV arr_of
using SC.is_concrete_set_category comp_def SC.set_img_Collect_ide_iff arr_of_def
by auto
+ lemma is_concrete_set_category:
+ shows "concrete_set_category comp setp UNIV arr_of"
+ ..
+
lemma bij_arr_of:
shows "bij_betw arr_of UNIV Univ"
using SC.bij_arr_of comp_def arr_of_def by presburger
lemma bij_elem_of:
shows "bij_betw elem_of Univ UNIV"
using SC.bij_elem_of comp_def elem_of_def by auto
end
-
+(*
+ interpretation RepleteSetCat: replete_setcat undefined .
+ *)
end
diff --git a/thys/Category3/SetCategory.thy b/thys/Category3/SetCategory.thy
--- a/thys/Category3/SetCategory.thy
+++ b/thys/Category3/SetCategory.thy
@@ -1,2676 +1,2680 @@
(* Title: SetCategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2016
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter SetCategory
theory SetCategory
imports Category Functor Subcategory
begin
text\<open>
This theory defines a locale \<open>set_category\<close> that axiomatizes the notion
``category of @{typ 'a}-sets and functions between them'' in the context of HOL.
A primary reason for doing this is to make it possible to prove results
(such as the Yoneda Lemma) that use such categories without having to commit to a
particular element type @{typ 'a} and without having the results depend on the
concrete details of a particular construction.
The axiomatization given here is categorical, in the sense that if categories
@{term S} and @{term S'} each interpret the \<open>set_category\<close> locale,
then a bijection between the sets of terminal objects of @{term S} and @{term S'}
extends to an isomorphism of @{term S} and @{term S'} as categories.
The axiomatization is based on the following idea: if, for some type @{typ 'a},
category @{term S} is the category of all @{typ 'a}-sets and functions between
them, then the elements of type @{typ 'a} are in bijective correspondence with
the terminal objects of category @{term S}. In addition, if @{term unity}
is an arbitrarily chosen terminal object of @{term S}, then for each object @{term a},
the hom-set @{term "hom unity a"} (\emph{i.e.} the set of ``points'' or
``global elements'' of @{term a}) is in bijective correspondence with a subset
of the terminal objects of @{term S}. By making a specific, but arbitrary,
choice of such a correspondence, we can then associate with each object @{term a}
of @{term S} a set @{term "set a"} that consists of all terminal objects @{term t}
that correspond to some point @{term x} of @{term a}. Each arrow @{term f}
then induces a function \<open>Fun f \<in> set (dom f) \<rightarrow> set (cod f)\<close>,
defined on terminal objects of @{term S} by passing to points of @{term "dom f"},
composing with @{term f}, then passing back from points of @{term "cod f"}
to terminal objects. Once we can associate a set with each object of @{term S}
and a function with each arrow, we can force @{term S} to be isomorphic to the
category of @{typ 'a}-sets by imposing suitable extensionality and completeness axioms.
\<close>
section "Some Lemmas about Restriction"
text\<open>
\sloppypar
The development of the \<open>set_category\<close> locale makes heavy use of the
theory @{theory "HOL-Library.FuncSet"}. However, in some cases, I found that
that theory did not provide results about restriction in the form that was
most useful to me. I used the following additional results in various places.
\<close>
(* This variant of FuncSet.restrict_ext is sometimes useful. *)
lemma restr_eqI:
assumes "A = A'" and "\<And>x. x \<in> A \<Longrightarrow> F x = F' x"
shows "restrict F A = restrict F' A'"
using assms by force
(* This rule avoided a long proof in at least one instance
where FuncSet.restrict_apply did not work.
*)
lemma restr_eqE [elim]:
assumes "restrict F A = restrict F' A" and "x \<in> A"
shows "F x = F' x"
using assms restrict_def by metis
(* This seems more useful than compose_eq in FuncSet. *)
lemma compose_eq' [simp]:
shows "compose A G F = restrict (G o F) A"
unfolding compose_def restrict_def by auto
section "Set Categories"
text\<open>
We first define the locale \<open>set_category_data\<close>, which sets out the basic
data and definitions for the \<open>set_category\<close> locale, without imposing any conditions
other than that @{term S} is a category and that @{term img} is a function defined
on the arrow type of @{term S}. The function @{term img} should be thought of
as a mapping that takes a point @{term "x \<in> hom unity a"} to a corresponding
terminal object @{term "img x"}. Eventually, assumptions will be introduced so
that this is in fact the case. The set of terminal objects of the category will
serve as abstract ``elements'' of sets; we will refer to the set of \emph{all}
terminal objects as the \emph{universe}.
\<close>
locale set_category_data = category S
for S :: "'s comp" (infixr "\<cdot>" 55)
and img :: "'s \<Rightarrow> 's"
begin
notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text\<open>
Call the set of all terminal objects of S the ``universe''.
\<close>
abbreviation Univ :: "'s set"
where "Univ \<equiv> Collect terminal"
text\<open>
Choose an arbitrary element of the universe and call it @{term unity}.
\<close>
definition unity :: 's
where "unity = (SOME t. terminal t)"
text\<open>
Each object @{term a} determines a subset @{term "set a"} of the universe,
consisting of all those terminal objects @{term t} such that @{term "t = img x"}
for some @{term "x \<in> hom unity a"}.
\<close>
definition set :: "'s \<Rightarrow> 's set"
where "set a = img ` hom unity a"
end
text\<open>
Next, we define a locale \<open>set_category_given_img\<close> that augments the
\<open>set_category_data\<close> locale with assumptions that serve to define
the notion of a set category with a chosen correspondence between points
and terminal objects. The assumptions require that the universe be nonempty
(so that the definition of @{term unity} makes sense), that the map
@{term img} is a locally injective map taking points to terminal objects,
that each terminal object @{term t} belongs to @{term "set t"},
that two objects of @{term S} are equal if they determine the same set,
that two parallel arrows of @{term S} are equal if they determine the same
function, and that for any objects @{term a} and @{term b} and function
@{term "F \<in> hom unity a \<rightarrow> hom unity b"} there is an arrow @{term "f \<in> hom a b"}
whose action under the composition of @{term S} coincides with the function @{term F}.
The parameter @{term setp} is a predicate that determines which subsets of the
universe are to be regarded as defining objects of the category. This parameter
has been introduced because most of the characteristic properties of a category
of sets and functions do not depend on there being an object corresponding to
\emph{every} subset of the universe, and we intend to consider in particular the
cases in which only finite subsets or only ``small'' subsets of the universe
determine objects. Accordingly, we assume that there is an object corresponding
to each subset of the universe that satisfies @{term setp}.
It is also necessary to assume some basic regularity properties of the predicate
@{term setp}; namely, that it holds for all subsets of the universe corresponding
to objects of @{term S}, and that it respects subset and union.
\<close>
locale set_category_given_img = set_category_data S img
for S :: "'s comp" (infixr "\<cdot>" 55)
and img :: "'s \<Rightarrow> 's"
and setp :: "'s set \<Rightarrow> bool" +
assumes setp_imp_subset_Univ: "setp A \<Longrightarrow> A \<subseteq> Univ"
and setp_set_ide: "ide a \<Longrightarrow> setp (set a)"
and setp_respects_subset: "A' \<subseteq> A \<Longrightarrow> setp A \<Longrightarrow> setp A'"
and setp_respects_union: "\<lbrakk> setp A; setp B \<rbrakk> \<Longrightarrow> setp (A \<union> B)"
and nonempty_Univ: "Univ \<noteq> {}"
and inj_img: "ide a \<Longrightarrow> inj_on img (hom unity a)"
and stable_img: "terminal t \<Longrightarrow> t \<in> img ` hom unity t"
and extensional_set: "\<lbrakk> ide a; ide b; set a = set b \<rbrakk> \<Longrightarrow> a = b"
and extensional_arr: "\<lbrakk> par f f'; \<And>x. \<guillemotleft>x : unity \<rightarrow> dom f\<guillemotright> \<Longrightarrow> f \<cdot> x = f' \<cdot> x \<rbrakk> \<Longrightarrow> f = f'"
and set_complete: "setp A \<Longrightarrow> \<exists>a. ide a \<and> set a = A"
and fun_complete_ax: "\<lbrakk> ide a; ide b; F \<in> hom unity a \<rightarrow> hom unity b \<rbrakk>
\<Longrightarrow> \<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : unity \<rightarrow> dom f\<guillemotright> \<longrightarrow> f \<cdot> x = F x)"
begin
lemma setp_singleton:
assumes "terminal a"
shows "setp {a}"
using assms
by (metis setp_set_ide Set.set_insert Un_upper1 insert_is_Un set_def
setp_respects_subset stable_img terminal_def)
lemma setp_empty:
shows "setp {}"
using setp_singleton setp_respects_subset nonempty_Univ by blast
lemma finite_imp_setp:
assumes "A \<subseteq> Univ" and "finite A"
shows "setp A"
using setp_empty setp_singleton setp_respects_union
by (metis assms(1-2) finite_subset_induct insert_is_Un mem_Collect_eq)
text\<open>
Each arrow @{term "f \<in> hom a b"} determines a function @{term "Fun f \<in> Univ \<rightarrow> Univ"},
by passing from @{term Univ} to @{term "hom a unity"}, composing with @{term f},
then passing back to @{term Univ}.
\<close>
definition Fun :: "'s \<Rightarrow> 's \<Rightarrow> 's"
where "Fun f = restrict (img o S f o inv_into (hom unity (dom f)) img) (set (dom f))"
- lemma comp_arr_point:
+ lemma comp_arr_point\<^sub>S\<^sub>C:
assumes "arr f" and "\<guillemotleft>x : unity \<rightarrow> dom f\<guillemotright>"
shows "f \<cdot> x = inv_into (hom unity (cod f)) img (Fun f (img x))"
proof -
have "\<guillemotleft>f \<cdot> x : unity \<rightarrow> cod f\<guillemotright>"
using assms by blast
thus ?thesis
using assms Fun_def inj_img set_def by simp
qed
text\<open>
Parallel arrows that determine the same function are equal.
\<close>
- lemma arr_eqI:
+ lemma arr_eqI\<^sub>S\<^sub>C:
assumes "par f f'" and "Fun f = Fun f'"
shows "f = f'"
- using assms comp_arr_point extensional_arr by metis
+ using assms comp_arr_point\<^sub>S\<^sub>C extensional_arr by metis
- lemma terminal_unity:
+ lemma terminal_unity\<^sub>S\<^sub>C:
shows "terminal unity"
using unity_def nonempty_Univ by (simp add: someI_ex)
lemma ide_unity [simp]:
shows "ide unity"
- using terminal_unity terminal_def by blast
+ using terminal_unity\<^sub>S\<^sub>C terminal_def by blast
lemma setp_set' [simp]:
assumes "ide a"
shows "setp (set a)"
using assms setp_set_ide by auto
lemma inj_on_set:
shows "inj_on set (Collect ide)"
using extensional_set by (intro inj_onI, auto)
text\<open>
The inverse of the map @{term set} is a map @{term mkIde} that takes each subset
of the universe to an identity of @{term[source=true] S}.
\<close>
definition mkIde :: "'s set \<Rightarrow> 's"
where "mkIde A = (if setp A then inv_into (Collect ide) set A else null)"
lemma mkIde_set [simp]:
assumes "ide a"
shows "mkIde (set a) = a"
by (simp add: assms inj_on_set mkIde_def)
lemma set_mkIde [simp]:
assumes "setp A"
shows "set (mkIde A) = A"
using assms mkIde_def set_complete someI_ex [of "\<lambda>a. a \<in> Collect ide \<and> set a = A"]
mkIde_set
by metis
lemma ide_mkIde [simp]:
assumes "setp A"
shows "ide (mkIde A)"
using assms mkIde_def mkIde_set set_complete by metis
lemma arr_mkIde [iff]:
shows "arr (mkIde A) \<longleftrightarrow> setp A"
using not_arr_null mkIde_def ide_mkIde by auto
lemma dom_mkIde [simp]:
assumes "setp A"
shows "dom (mkIde A) = mkIde A"
using assms ide_mkIde by simp
lemma cod_mkIde [simp]:
assumes "setp A"
shows "cod (mkIde A) = mkIde A"
using assms ide_mkIde by simp
text\<open>
Each arrow @{term f} determines an extensional function from
@{term "set (dom f)"} to @{term "set (cod f)"}.
\<close>
lemma Fun_mapsto:
assumes "arr f"
shows "Fun f \<in> extensional (set (dom f)) \<inter> (set (dom f) \<rightarrow> set (cod f))"
proof
show "Fun f \<in> extensional (set (dom f))" using Fun_def by fastforce
show "Fun f \<in> set (dom f) \<rightarrow> set (cod f)"
proof
fix t
assume t: "t \<in> set (dom f)"
have "Fun f t = img (f \<cdot> inv_into (hom unity (dom f)) img t)"
using assms t Fun_def comp_def by simp
moreover have "... \<in> set (cod f)"
using assms t set_def inv_into_into [of t img "hom unity (dom f)"] by blast
ultimately show "Fun f t \<in> set (cod f)" by auto
qed
qed
text\<open>
Identities of @{term[source=true] S} correspond to restrictions of the identity function.
\<close>
lemma Fun_ide:
assumes "ide a"
shows "Fun a = restrict (\<lambda>x. x) (set a)"
using assms Fun_def inj_img set_def comp_cod_arr by fastforce
lemma Fun_mkIde [simp]:
assumes "setp A"
shows "Fun (mkIde A) = restrict (\<lambda>x. x) A"
using assms ide_mkIde set_mkIde Fun_ide by simp
text\<open>
Composition in @{term S} corresponds to extensional function composition.
\<close>
lemma Fun_comp [simp]:
assumes "seq g f"
shows "Fun (g \<cdot> f) = restrict (Fun g o Fun f) (set (dom f))"
proof -
have "restrict (img o S (g \<cdot> f) o (inv_into (hom unity (dom (g \<cdot> f))) img))
(set (dom (g \<cdot> f)))
= restrict (Fun g o Fun f) (set (dom f))"
proof -
let ?img' = "\<lambda>a. \<lambda>t. inv_into (hom unity a) img t"
have 1: "set (dom (g \<cdot> f)) = set (dom f)"
using assms by auto
moreover have "\<And>t. t \<in> set (dom (g \<cdot> f)) \<Longrightarrow>
(img o S (g \<cdot> f) o ?img' (dom (g \<cdot> f))) t = (Fun g o Fun f) t"
proof -
fix t
assume "t \<in> set (dom (g \<cdot> f))"
hence t: "t \<in> set (dom f)" by (simp add: 1)
have "(img o S (g \<cdot> f) o ?img' (dom (g \<cdot> f))) t = img (g \<cdot> f \<cdot> ?img' (dom f) t)"
using assms dom_comp comp_assoc by simp
also have "... = img (g \<cdot> ?img' (dom g) (Fun f t))"
proof -
have "\<And>a x. x \<in> hom unity a \<Longrightarrow> ?img' a (img x) = x"
using assms inj_img ide_cod inv_into_f_eq
by (metis arrI in_homE mem_Collect_eq)
thus ?thesis
- using assms t Fun_def set_def comp_arr_point by auto
+ using assms t Fun_def set_def comp_arr_point\<^sub>S\<^sub>C by auto
qed
also have "... = Fun g (Fun f t)"
proof -
have "Fun f t \<in> img ` hom unity (cod f)"
using assms t Fun_mapsto set_def by fast
thus ?thesis
using assms by (auto simp add: set_def Fun_def)
qed
finally show "(img o S (g \<cdot> f) o ?img' (dom (g \<cdot> f))) t = (Fun g o Fun f) t"
by auto
qed
ultimately show ?thesis by auto
qed
thus ?thesis using Fun_def by auto
qed
text\<open>
The constructor @{term mkArr} is used to obtain an arrow given subsets
@{term A} and @{term B} of the universe and a function @{term "F \<in> A \<rightarrow> B"}.
\<close>
definition mkArr :: "'s set \<Rightarrow> 's set \<Rightarrow> ('s \<Rightarrow> 's) \<Rightarrow> 's"
where "mkArr A B F = (if setp A \<and> setp B \<and> F \<in> A \<rightarrow> B
then (THE f. f \<in> hom (mkIde A) (mkIde B) \<and> Fun f = restrict F A)
else null)"
text\<open>
Each function @{term "F \<in> set a \<rightarrow> set b"} determines a unique arrow @{term "f \<in> hom a b"},
such that @{term "Fun f"} is the restriction of @{term F} to @{term "set a"}.
\<close>
lemma fun_complete:
assumes "ide a" and "ide b" and "F \<in> set a \<rightarrow> set b"
shows "\<exists>!f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> Fun f = restrict F (set a)"
proof -
let ?P = "\<lambda>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> Fun f = restrict F (set a)"
show "\<exists>!f. ?P f"
proof
have "\<exists>f. ?P f"
proof -
let ?F' = "\<lambda>x. inv_into (hom unity b) img (F (img x))"
have "?F' \<in> hom unity a \<rightarrow> hom unity b"
proof
fix x
assume x: "x \<in> hom unity a"
have "F (img x) \<in> set b" using assms(3) x set_def by auto
thus "inv_into (hom unity b) img (F (img x)) \<in> hom unity b"
using assms setp_set_ide inj_img set_def by auto
qed
hence "\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<longrightarrow> f \<cdot> x = ?F' x)"
using assms fun_complete_ax [of a b] by force
from this obtain f where f: "\<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<longrightarrow> f \<cdot> x = ?F' x)"
by blast
let ?img' = "\<lambda>a. \<lambda>t. inv_into (hom unity a) img t"
have "Fun f = restrict F (set a)"
proof (unfold Fun_def, intro restr_eqI)
show "set (dom f) = set a" using f by auto
show "\<And>t. t \<in> set (dom f) \<Longrightarrow> (img \<circ> S f \<circ> ?img' (dom f)) t = F t"
proof -
fix t
assume t: "t \<in> set (dom f)"
have "(img \<circ> S f \<circ> ?img' (dom f)) t = img (f \<cdot> ?img' (dom f) t)"
by simp
also have "... = img (?F' (?img' (dom f) t))"
by (metis f in_homE inv_into_into set_def mem_Collect_eq t)
also have "... = img (?img' (cod f) (F t))"
using f t set_def inj_img by auto
also have "... = F t"
proof -
have "F t \<in> set (cod f)"
using assms f t by auto
thus ?thesis
using f t set_def inj_img by auto
qed
finally show "(img \<circ> S f \<circ> ?img' (dom f)) t = F t" by auto
qed
qed
thus ?thesis using f by blast
qed
thus F: "?P (SOME f. ?P f)" using someI_ex [of ?P] by fast
show "\<And>f'. ?P f' \<Longrightarrow> f' = (SOME f. ?P f)"
- using F arr_eqI
+ using F arr_eqI\<^sub>S\<^sub>C
by (metis (no_types, lifting) in_homE)
qed
qed
lemma mkArr_in_hom:
assumes "setp A" and "setp B" and "F \<in> A \<rightarrow> B"
shows "\<guillemotleft>mkArr A B F : mkIde A \<rightarrow> mkIde B\<guillemotright>"
using assms mkArr_def fun_complete [of "mkIde A" "mkIde B" F] ide_mkIde set_mkIde
theI' [of "\<lambda>f. f \<in> hom (mkIde A) (mkIde B) \<and> Fun f = restrict F A"]
setp_imp_subset_Univ
by simp
text\<open>
The ``only if'' direction of the next lemma can be achieved only if there exists a
non-arrow element of type @{typ 's}, which can be used as the value of @{term "mkArr A B F"}
in cases where @{term "F \<notin> A \<rightarrow> B"}. Nevertheless, it is essential to have this,
because without the ``only if'' direction, we can't derive any useful
consequences from an assumption of the form @{term "arr (mkArr A B F)"};
instead we have to obtain @{term "F \<in> A \<rightarrow> B"} some other way.
This is is usually highly inconvenient and it makes the theory very weak and almost
unusable in practice. The observation that having a non-arrow value of type @{typ 's}
solves this problem is ultimately what led me to incorporate @{term null} first into the
definition of the \<open>set_category\<close> locale and then, ultimately, into the definition
of the \<open>category\<close> locale. I believe this idea is critical to the usability of the
entire development.
\<close>
lemma arr_mkArr [iff]:
shows "arr (mkArr A B F) \<longleftrightarrow> setp A \<and> setp B \<and> F \<in> A \<rightarrow> B"
proof
show "arr (mkArr A B F) \<Longrightarrow> setp A \<and> setp B \<and> F \<in> A \<rightarrow> B"
using mkArr_def not_arr_null ex_un_null someI_ex [of "\<lambda>f. \<not>arr f"] setp_imp_subset_Univ
by metis
show "setp A \<and> setp B \<and> F \<in> A \<rightarrow> B \<Longrightarrow> arr (mkArr A B F)"
using mkArr_in_hom by auto
qed
lemma arr_mkArrI [intro]:
assumes "setp A" and "setp B" and "F \<in> A \<rightarrow> B"
shows "arr (mkArr A B F)"
using assms arr_mkArr by blast
lemma Fun_mkArr':
assumes "arr (mkArr A B F)"
shows "\<guillemotleft>mkArr A B F : mkIde A \<rightarrow> mkIde B\<guillemotright>"
and "Fun (mkArr A B F) = restrict F A"
proof -
have 1: "setp A \<and> setp B \<and> F \<in> A \<rightarrow> B" using assms by fast
have 2: "mkArr A B F \<in> hom (mkIde A) (mkIde B) \<and>
Fun (mkArr A B F) = restrict F (set (mkIde A))"
proof -
have "\<exists>!f. f \<in> hom (mkIde A) (mkIde B) \<and> Fun f = restrict F (set (mkIde A))"
using 1 fun_complete [of "mkIde A" "mkIde B" F] ide_mkIde set_mkIde by simp
thus ?thesis using 1 mkArr_def theI' set_mkIde by simp
qed
show "\<guillemotleft>mkArr A B F : mkIde A \<rightarrow> mkIde B\<guillemotright>" using 1 2 by auto
show "Fun (mkArr A B F) = restrict F A" using 1 2 set_mkIde by auto
qed
lemma mkArr_Fun:
assumes "arr f"
shows "mkArr (set (dom f)) (set (cod f)) (Fun f) = f"
proof -
have 1: "setp (set (dom f)) \<and> setp (set (cod f)) \<and> ide (dom f) \<and> ide (cod f) \<and>
Fun f \<in> extensional (set (dom f)) \<inter> (set (dom f) \<rightarrow> set (cod f))"
using Fun_mapsto assms ide_cod ide_dom setp_set' by presburger
hence "\<exists>!f'. f' \<in> hom (dom f) (cod f) \<and> Fun f' = restrict (Fun f) (set (dom f))"
using fun_complete by force
moreover have "f \<in> hom (dom f) (cod f) \<and> Fun f = restrict (Fun f) (set (dom f))"
using assms 1 extensional_restrict by force
ultimately have "f = (THE f'. f' \<in> hom (dom f) (cod f) \<and>
Fun f' = restrict (Fun f) (set (dom f)))"
using theI' [of "\<lambda>f'. f' \<in> hom (dom f) (cod f) \<and> Fun f' = restrict (Fun f) (set (dom f))"]
by blast
also have "... = mkArr (set (dom f)) (set (cod f)) (Fun f)"
using assms 1 mkArr_def mkIde_set by simp
finally show ?thesis by auto
qed
lemma dom_mkArr [simp]:
assumes "arr (mkArr A B F)"
shows "dom (mkArr A B F) = mkIde A"
using assms Fun_mkArr' by auto
lemma cod_mkArr [simp]:
assumes "arr (mkArr A B F)"
shows "cod (mkArr A B F) = mkIde B"
using assms Fun_mkArr' by auto
lemma Fun_mkArr [simp]:
assumes "arr (mkArr A B F)"
shows "Fun (mkArr A B F) = restrict F A"
using assms Fun_mkArr' by auto
text\<open>
The following provides the basic technique for showing that arrows
constructed using @{term mkArr} are equal.
\<close>
lemma mkArr_eqI [intro]:
assumes "arr (mkArr A B F)"
and "A = A'" and "B = B'" and "\<And>x. x \<in> A \<Longrightarrow> F x = F' x"
shows "mkArr A B F = mkArr A' B' F'"
using assms Fun_mkArr
- by (intro arr_eqI, auto simp add: Pi_iff)
+ by (intro arr_eqI\<^sub>S\<^sub>C, auto simp add: Pi_iff)
text\<open>
This version avoids trivial proof obligations when the domain and codomain
sets are identical from the context.
\<close>
lemma mkArr_eqI' [intro]:
assumes "arr (mkArr A B F)" and "\<And>x. x \<in> A \<Longrightarrow> F x = F' x"
shows "mkArr A B F = mkArr A B F'"
using assms mkArr_eqI by simp
lemma mkArr_restrict_eq:
assumes "arr (mkArr A B F)"
shows "mkArr A B (restrict F A) = mkArr A B F"
using assms by (intro mkArr_eqI', auto)
lemma mkArr_restrict_eq':
assumes "arr (mkArr A B (restrict F A))"
shows "mkArr A B (restrict F A) = mkArr A B F"
using assms by (intro mkArr_eqI', auto)
lemma mkIde_as_mkArr [simp]:
assumes "setp A"
shows "mkArr A A (\<lambda>x. x) = mkIde A"
using assms arr_mkIde dom_mkIde cod_mkIde Fun_mkIde
- by (intro arr_eqI, auto)
+ by (intro arr_eqI\<^sub>S\<^sub>C, auto)
lemma comp_mkArr:
assumes "arr (mkArr A B F)" and "arr (mkArr B C G)"
shows "mkArr B C G \<cdot> mkArr A B F = mkArr A C (G \<circ> F)"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 1: "seq (mkArr B C G) (mkArr A B F)" using assms by force
have 2: "G o F \<in> A \<rightarrow> C" using assms by auto
show "par (mkArr B C G \<cdot> mkArr A B F) (mkArr A C (G \<circ> F))"
using assms 1 2
by (intro conjI) simp_all
show "Fun (mkArr B C G \<cdot> mkArr A B F) = Fun (mkArr A C (G \<circ> F))"
using 1 2 by fastforce
qed
text\<open>
The locale assumption @{prop stable_img} forces @{term "t \<in> set t"} in case
@{term t} is a terminal object. This is very convenient, as it results in the
characterization of terminal objects as identities @{term t} for which
@{term "set t = {t}"}. However, it is not absolutely necessary to have this.
The following weaker characterization of terminal objects can be proved without
the @{prop stable_img} assumption.
\<close>
lemma terminal_char1:
shows "terminal t \<longleftrightarrow> ide t \<and> (\<exists>!x. x \<in> set t)"
proof -
have "terminal t \<Longrightarrow> ide t \<and> (\<exists>!x. x \<in> set t)"
proof -
assume t: "terminal t"
have "ide t" using t terminal_def by auto
moreover have "\<exists>!x. x \<in> set t"
proof -
have "\<exists>!x. x \<in> hom unity t"
- using t terminal_unity terminal_def by auto
+ using t terminal_unity\<^sub>S\<^sub>C terminal_def by auto
thus ?thesis using set_def by auto
qed
ultimately show "ide t \<and> (\<exists>!x. x \<in> set t)" by auto
qed
moreover have "ide t \<and> (\<exists>!x. x \<in> set t) \<Longrightarrow> terminal t"
proof -
assume t: "ide t \<and> (\<exists>!x. x \<in> set t)"
from this obtain t' where "set t = {t'}" by blast
hence t': "set t = {t'} \<and> setp {t'} \<and> t = mkIde {t'}"
using t setp_set_ide mkIde_set by metis
show "terminal t"
proof
show "ide t" using t by simp
show "\<And>a. ide a \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a \<rightarrow> t\<guillemotright>"
proof -
fix a
assume a: "ide a"
show "\<exists>!f. \<guillemotleft>f : a \<rightarrow> t\<guillemotright>"
proof
show 1: "\<guillemotleft>mkArr (set a) {t'} (\<lambda>x. t') : a \<rightarrow> t\<guillemotright>"
using a t t' mkArr_in_hom
by (metis Pi_I' mkIde_set setp_set_ide singletonD)
show "\<And>f. \<guillemotleft>f : a \<rightarrow> t\<guillemotright> \<Longrightarrow> f = mkArr (set a) {t'} (\<lambda>x. t')"
proof -
fix f
assume f: "\<guillemotleft>f : a \<rightarrow> t\<guillemotright>"
show "f = mkArr (set a) {t'} (\<lambda>x. t')"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
show 1: "par f (mkArr (set a) {t'} (\<lambda>x. t'))" using 1 f in_homE by metis
show "Fun f = Fun (mkArr (set a) {t'} (\<lambda>x. t'))"
proof -
have "Fun (mkArr (set a) {t'} (\<lambda>x. t')) = (\<lambda>x\<in>set a. t')"
using 1 Fun_mkArr by simp
also have "... = Fun f"
proof -
have "\<And>x. x \<in> set a \<Longrightarrow> Fun f x = t'"
using f t' Fun_def mkArr_Fun arr_mkArr
by (metis PiE in_homE singletonD)
moreover have "\<And>x. x \<notin> set a \<Longrightarrow> Fun f x = undefined"
using f Fun_def by auto
ultimately show ?thesis by auto
qed
finally show ?thesis by force
qed
qed
qed
qed
qed
qed
qed
ultimately show ?thesis by blast
qed
text\<open>
As stated above, in the presence of the @{prop stable_img} assumption we have
the following stronger characterization of terminal objects.
\<close>
lemma terminal_char2:
shows "terminal t \<longleftrightarrow> ide t \<and> set t = {t}"
proof
assume t: "terminal t"
show "ide t \<and> set t = {t}"
proof
show "ide t" using t terminal_char1 by auto
show "set t = {t}"
proof -
- have "\<exists>!x. x \<in> hom unity t" using t terminal_def terminal_unity by force
+ have "\<exists>!x. x \<in> hom unity t" using t terminal_def terminal_unity\<^sub>S\<^sub>C by force
moreover have "t \<in> img ` hom unity t" using t stable_img set_def by simp
ultimately show ?thesis using set_def by auto
qed
qed
next
assume "ide t \<and> set t = {t}"
thus "terminal t" using terminal_char1 by force
qed
end
text\<open>
At last, we define the \<open>set_category\<close> locale by existentially quantifying
out the choice of a particular @{term img} map. We need to know that such a map
exists, but it does not matter which one we choose.
\<close>
locale set_category = category S
for S :: "'s comp" (infixr "\<cdot>" 55)
and setp :: "'s set \<Rightarrow> bool" +
assumes ex_img: "\<exists>img. set_category_given_img S img setp"
begin
notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
definition some_img
where "some_img = (SOME img. set_category_given_img S img setp)"
sublocale set_category_given_img S some_img setp
proof -
have "\<exists>img. set_category_given_img S img setp" using ex_img by auto
thus "set_category_given_img S some_img setp"
using someI_ex [of "\<lambda>img. set_category_given_img S img setp"] some_img_def
by metis
qed
end
text\<open>We call a set category \emph{replete} if there is an object corresponding to
every subset of the universe.
\<close>
locale replete_set_category =
category S +
set_category S \<open>\<lambda>A. A \<subseteq> Collect terminal\<close>
for S :: "'s comp" (infixr "\<cdot>" 55)
begin
abbreviation setp
where "setp \<equiv> \<lambda>A. A \<subseteq> Univ"
+ lemma is_set_category:
+ shows "set_category S (\<lambda>A. A \<subseteq> Collect terminal)"
+ ..
+
end
context set_category
begin
text\<open>
The arbitrary choice of @{term img} induces a system of arrows corresponding
to inclusions of subsets.
\<close>
definition incl :: "'s \<Rightarrow> bool"
where "incl f = (arr f \<and> set (dom f) \<subseteq> set (cod f) \<and>
f = mkArr (set (dom f)) (set (cod f)) (\<lambda>x. x))"
lemma Fun_incl:
assumes "incl f"
shows "Fun f = (\<lambda>x \<in> set (dom f). x)"
using assms incl_def by (metis Fun_mkArr)
lemma ex_incl_iff_subset:
assumes "ide a" and "ide b"
shows "(\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> incl f) \<longleftrightarrow> set a \<subseteq> set b"
proof
show "\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> incl f \<Longrightarrow> set a \<subseteq> set b"
using incl_def by auto
show "set a \<subseteq> set b \<Longrightarrow> \<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> incl f"
proof
assume 1: "set a \<subseteq> set b"
show "\<guillemotleft>mkArr (set a) (set b) (\<lambda>x. x) : a \<rightarrow> b\<guillemotright> \<and> incl (mkArr (set a) (set b) (\<lambda>x. x))"
proof
show "\<guillemotleft>mkArr (set a) (set b) (\<lambda>x. x) : a \<rightarrow> b\<guillemotright>"
by (metis 1 assms image_ident image_subset_iff_funcset mkIde_set
mkArr_in_hom setp_set_ide)
thus "incl (mkArr (set a) (set b) (\<lambda>x. x))"
using 1 incl_def by force
qed
qed
qed
end
section "Categoricity"
text\<open>
In this section we show that the \<open>set_category\<close> locale completely characterizes
the structure of its interpretations as categories, in the sense that for any two
interpretations @{term S} and @{term S'}, a @{term setp}-respecting bijection
between the universe of @{term S} and the universe of @{term S'} extends
to an isomorphism of @{term S} and @{term S'}.
\<close>
locale two_set_categories_bij_betw_Univ =
S: set_category S setp +
S': set_category S' setp'
for S :: "'s comp" (infixr "\<cdot>" 55)
and setp :: "'s set \<Rightarrow> bool"
and S' :: "'t comp" (infixr "\<cdot>\<acute>" 55)
and setp' :: "'t set \<Rightarrow> bool"
and \<phi> :: "'s \<Rightarrow> 't" +
assumes bij_\<phi>: "bij_betw \<phi> S.Univ S'.Univ"
and \<phi>_respects_setp: "A \<subseteq> S.Univ \<Longrightarrow> setp' (\<phi> ` A) \<longleftrightarrow> setp A"
begin
notation S.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation S'.in_hom ("\<guillemotleft>_ : _ \<rightarrow>'' _\<guillemotright>")
abbreviation \<psi>
where "\<psi> \<equiv> inv_into S.Univ \<phi>"
lemma \<psi>_\<phi>:
assumes "t \<in> S.Univ"
shows "\<psi> (\<phi> t) = t"
using assms bij_\<phi> bij_betw_inv_into_left by metis
lemma \<phi>_\<psi>:
assumes "t' \<in> S'.Univ"
shows "\<phi> (\<psi> t') = t'"
using assms bij_\<phi> bij_betw_inv_into_right by metis
lemma \<psi>_img_\<phi>_img:
assumes "A \<subseteq> S.Univ"
shows "\<psi> ` \<phi> ` A = A"
using assms bij_\<phi> by (simp add: bij_betw_def)
lemma \<phi>_img_\<psi>_img:
assumes "A' \<subseteq> S'.Univ"
shows "\<phi> ` \<psi> ` A' = A'"
using assms bij_\<phi> by (simp add: bij_betw_def image_inv_into_cancel)
text\<open>
We define the object map @{term \<Phi>o} of a functor from @{term[source=true] S}
to @{term[source=true] S'}.
\<close>
definition \<Phi>o
where "\<Phi>o = (\<lambda>a \<in> Collect S.ide. S'.mkIde (\<phi> ` S.set a))"
lemma set_\<Phi>o:
assumes "S.ide a"
shows "S'.set (\<Phi>o a) = \<phi> ` S.set a"
by (simp add: S.setp_imp_subset_Univ \<Phi>o_def \<phi>_respects_setp assms)
lemma \<Phi>o_preserves_ide:
assumes "S.ide a"
shows "S'.ide (\<Phi>o a)"
using assms S'.ide_mkIde bij_\<phi> bij_betw_def image_mono restrict_apply' S.setp_set'
\<phi>_respects_setp S.setp_imp_subset_Univ
unfolding \<Phi>o_def
by simp
text\<open>
The map @{term \<Phi>a} assigns to each arrow @{term f} of @{term[source=true] S} the function on
the universe of @{term[source=true] S'} that is the same as the function induced by @{term f}
on the universe of @{term[source=true] S}, up to the bijection @{term \<phi>} between the two
universes.
\<close>
definition \<Phi>a
where "\<Phi>a = (\<lambda>f. \<lambda>x' \<in> \<phi> ` S.set (S.dom f). \<phi> (S.Fun f (\<psi> x')))"
lemma \<Phi>a_mapsto:
assumes "S.arr f"
shows "\<Phi>a f \<in> S'.set (\<Phi>o (S.dom f)) \<rightarrow> S'.set (\<Phi>o (S.cod f))"
proof -
have "\<Phi>a f \<in> \<phi> ` S.set (S.dom f) \<rightarrow> \<phi> ` S.set (S.cod f)"
proof
fix x
assume x: "x \<in> \<phi> ` S.set (S.dom f)"
have "\<psi> x \<in> S.set (S.dom f)"
using assms x \<psi>_img_\<phi>_img [of "S.set (S.dom f)"] S.setp_imp_subset_Univ by auto
hence "S.Fun f (\<psi> x) \<in> S.set (S.cod f)" using assms S.Fun_mapsto by auto
hence "\<phi> (S.Fun f (\<psi> x)) \<in> \<phi> ` S.set (S.cod f)" by simp
thus "\<Phi>a f x \<in> \<phi> ` S.set (S.cod f)" using x \<Phi>a_def by auto
qed
thus ?thesis using assms set_\<Phi>o \<Phi>o_preserves_ide by auto
qed
text\<open>
The map @{term \<Phi>a} takes composition of arrows to extensional
composition of functions.
\<close>
lemma \<Phi>a_comp:
assumes gf: "S.seq g f"
shows "\<Phi>a (g \<cdot> f) = restrict (\<Phi>a g o \<Phi>a f) (S'.set (\<Phi>o (S.dom f)))"
proof -
have "\<Phi>a (g \<cdot> f) = (\<lambda>x' \<in> \<phi> ` S.set (S.dom f). \<phi> (S.Fun (S g f) (\<psi> x')))"
using gf \<Phi>a_def by auto
also have "... = (\<lambda>x' \<in> \<phi> ` S.set (S.dom f).
\<phi> (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\<psi> x')))"
using gf set_\<Phi>o S.Fun_comp by simp
also have "... = restrict (\<Phi>a g o \<Phi>a f) (S'.set (\<Phi>o (S.dom f)))"
proof -
have "\<And>x'. x' \<in> \<phi> ` S.set (S.dom f)
\<Longrightarrow> \<phi> (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\<psi> x')) = \<Phi>a g (\<Phi>a f x')"
proof -
fix x'
assume X': "x' \<in> \<phi> ` S.set (S.dom f)"
hence 1: "\<psi> x' \<in> S.set (S.dom f)"
using gf \<psi>_img_\<phi>_img S.setp_imp_subset_Univ S.ide_dom S.setp_set_ide
by blast
hence "\<phi> (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\<psi> x'))
= \<phi> (S.Fun g (S.Fun f (\<psi> x')))"
using restrict_apply by auto
also have "... = \<phi> (S.Fun g (\<psi> (\<phi> (S.Fun f (\<psi> x')))))"
proof -
have "S.Fun f (\<psi> x') \<in> S.set (S.cod f)"
using gf 1 S.Fun_mapsto by fast
hence "\<psi> (\<phi> (S.Fun f (\<psi> x'))) = S.Fun f (\<psi> x')"
using assms bij_\<phi> S.setp_imp_subset_Univ bij_betw_def inv_into_f_f subsetCE
S.ide_cod S.setp_set_ide
by (metis S.seqE)
thus ?thesis by auto
qed
also have "... = \<Phi>a g (\<Phi>a f x')"
proof -
have "\<Phi>a f x' \<in> \<phi> ` S.set (S.cod f)"
using gf S.ide_dom S.ide_cod X' \<Phi>a_mapsto [of f] set_\<Phi>o [of "S.dom f"]
set_\<Phi>o [of "S.cod f"]
by blast
thus ?thesis using gf X' \<Phi>a_def by auto
qed
finally show "\<phi> (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\<psi> x')) =
\<Phi>a g (\<Phi>a f x')"
by auto
qed
thus ?thesis using assms set_\<Phi>o by fastforce
qed
finally show ?thesis by auto
qed
text\<open>
Finally, we use @{term \<Phi>o} and @{term \<Phi>a} to define a functor @{term \<Phi>}.
\<close>
definition \<Phi>
where "\<Phi> f = (if S.arr f then
S'.mkArr (S'.set (\<Phi>o (S.dom f))) (S'.set (\<Phi>o (S.cod f))) (\<Phi>a f)
else S'.null)"
lemma \<Phi>_in_hom:
assumes "S.arr f"
shows "\<Phi> f \<in> S'.hom (\<Phi>o (S.dom f)) (\<Phi>o (S.cod f))"
proof -
have "\<guillemotleft>\<Phi> f : S'.dom (\<Phi> f) \<rightarrow>' S'.cod (\<Phi> f)\<guillemotright>"
using assms \<Phi>_def \<Phi>a_mapsto \<Phi>o_preserves_ide
by (intro S'.in_homI) auto
thus ?thesis
using assms \<Phi>_def \<Phi>a_mapsto \<Phi>o_preserves_ide by auto
qed
lemma \<Phi>_ide [simp]:
assumes "S.ide a"
shows "\<Phi> a = \<Phi>o a"
proof -
have "\<Phi> a = S'.mkArr (S'.set (\<Phi>o a)) (S'.set (\<Phi>o a)) (\<lambda>x'. x')"
proof -
have "\<guillemotleft>\<Phi> a : \<Phi>o a \<rightarrow>' \<Phi>o a\<guillemotright>"
using assms \<Phi>_in_hom S.ide_in_hom by fastforce
moreover have "\<Phi>a a = restrict (\<lambda>x'. x') (S'.set (\<Phi>o a))"
proof -
have "\<Phi>a a = (\<lambda>x' \<in> \<phi> ` S.set a. \<phi> (S.Fun a (\<psi> x')))"
using assms \<Phi>a_def restrict_apply by auto
also have "... = (\<lambda>x' \<in> S'.set (\<Phi>o a). \<phi> (\<psi> x'))"
proof -
have "S.Fun a = (\<lambda>x \<in> S.set a. x)"
using assms S.Fun_ide by auto
moreover have "\<And>x'. x' \<in> \<phi> ` S.set a \<Longrightarrow> \<psi> x' \<in> S.set a"
using assms bij_\<phi> S.setp_imp_subset_Univ image_iff S.setp_set_ide
by (metis \<psi>_img_\<phi>_img)
ultimately show ?thesis
using assms set_\<Phi>o by auto
qed
also have "... = restrict (\<lambda>x'. x') (S'.set (\<Phi>o a))"
using assms S'.setp_imp_subset_Univ S'.setp_set_ide \<Phi>o_preserves_ide \<phi>_\<psi>
by (meson restr_eqI subsetCE)
ultimately show ?thesis by auto
qed
ultimately show ?thesis
using assms \<Phi>_def \<Phi>o_preserves_ide S'.mkArr_restrict_eq'
by (metis S'.arrI S.ide_char)
qed
thus ?thesis
using assms S'.mkIde_as_mkArr \<Phi>o_preserves_ide \<Phi>_in_hom S'.mkIde_set
by simp
qed
lemma set_dom_\<Phi>:
assumes "S.arr f"
shows "S'.set (S'.dom (\<Phi> f)) = \<phi> ` (S.set (S.dom f))"
using assms S.ide_dom \<Phi>_in_hom \<Phi>_ide set_\<Phi>o by fastforce
lemma \<Phi>_comp:
assumes "S.seq g f"
shows "\<Phi> (g \<cdot> f) = \<Phi> g \<cdot>\<acute> \<Phi> f"
proof -
have "\<Phi> (g \<cdot> f) = S'.mkArr (S'.set (\<Phi>o (S.dom f))) (S'.set (\<Phi>o (S.cod g))) (\<Phi>a (S g f))"
using \<Phi>_def assms by auto
also have "... = S'.mkArr (S'.set (\<Phi>o (S.dom f))) (S'.set (\<Phi>o (S.cod g)))
(restrict (\<Phi>a g o \<Phi>a f) (S'.set (\<Phi>o (S.dom f))))"
using assms \<Phi>a_comp set_\<Phi>o by force
also have "... = S'.mkArr (S'.set (\<Phi>o (S.dom f))) (S'.set (\<Phi>o (S.cod g))) (\<Phi>a g o \<Phi>a f)"
by (metis S'.mkArr_restrict_eq' \<Phi>_in_hom assms calculation S'.in_homE mem_Collect_eq)
also have "... = S' (S'.mkArr (S'.set (\<Phi>o (S.dom g))) (S'.set (\<Phi>o (S.cod g))) (\<Phi>a g))
(S'.mkArr (S'.set (\<Phi>o (S.dom f))) (S'.set (\<Phi>o (S.cod f))) (\<Phi>a f))"
proof -
have "S'.arr (S'.mkArr (S'.set (\<Phi>o (S.dom f))) (S'.set (\<Phi>o (S.cod f))) (\<Phi>a f))"
using assms \<Phi>a_mapsto set_\<Phi>o S.ide_dom S.ide_cod \<Phi>o_preserves_ide
S'.arr_mkArr S'.setp_imp_subset_Univ S'.setp_set_ide S.seqE
by metis
moreover have "S'.arr (S'.mkArr (S'.set (\<Phi>o (S.dom g))) (S'.set (\<Phi>o (S.cod g)))
(\<Phi>a g))"
using assms \<Phi>a_mapsto set_\<Phi>o S.ide_dom S.ide_cod \<Phi>o_preserves_ide S'.arr_mkArr
S'.setp_imp_subset_Univ S'.setp_set_ide S.seqE
by metis
ultimately show ?thesis using assms S'.comp_mkArr by auto
qed
also have "... = \<Phi> g \<cdot>\<acute> \<Phi> f" using assms \<Phi>_def by force
finally show ?thesis by fast
qed
interpretation \<Phi>: "functor" S S' \<Phi>
apply unfold_locales
using \<Phi>_def
apply simp
using \<Phi>_in_hom \<Phi>_comp
by auto
lemma \<Phi>_is_functor:
shows "functor S S' \<Phi>" ..
lemma Fun_\<Phi>:
assumes "S.arr f" and "x \<in> S.set (S.dom f)"
shows "S'.Fun (\<Phi> f) (\<phi> x) = \<Phi>a f (\<phi> x)"
using assms \<Phi>_def \<Phi>.preserves_arr set_\<Phi>o by auto
lemma \<Phi>_acts_elementwise:
assumes "S.ide a"
shows "S'.set (\<Phi> a) = \<Phi> ` S.set a"
proof
have 0: "S'.set (\<Phi> a) = \<phi> ` S.set a"
using assms \<Phi>_ide set_\<Phi>o by simp
have 1: "\<And>x. x \<in> S.set a \<Longrightarrow> \<Phi> x = \<phi> x"
proof -
fix x
assume x: "x \<in> S.set a"
have 1: "S.terminal x" using assms x S.setp_imp_subset_Univ S.setp_set_ide by blast
hence 2: "S'.terminal (\<phi> x)"
by (metis CollectD CollectI bij_\<phi> bij_betw_def image_iff)
have "\<Phi> x = \<Phi>o x"
using assms x 1 \<Phi>_ide S.terminal_def by auto
also have "... = \<phi> x"
proof -
have "\<Phi>o x = S'.mkIde (\<phi> ` S.set x)"
using assms 1 x \<Phi>o_def S.terminal_def by auto
moreover have "S'.mkIde (\<phi> ` S.set x) = \<phi> x"
using assms x 1 2 S.terminal_char2 S'.terminal_char2 S'.mkIde_set bij_\<phi>
by (metis (no_types, lifting) empty_is_image image_insert)
ultimately show ?thesis by auto
qed
finally show "\<Phi> x = \<phi> x" by auto
qed
show "S'.set (\<Phi> a) \<subseteq> \<Phi> ` S.set a" using 0 1 by force
show "\<Phi> ` S.set a \<subseteq> S'.set (\<Phi> a)" using 0 1 by force
qed
lemma \<Phi>_preserves_incl:
assumes "S.incl m"
shows "S'.incl (\<Phi> m)"
proof -
have 1: "S.arr m \<and> S.set (S.dom m) \<subseteq> S.set (S.cod m) \<and>
m = S.mkArr (S.set (S.dom m)) (S.set (S.cod m)) (\<lambda>x. x)"
using assms S.incl_def by blast
have "S'.arr (\<Phi> m)" using 1 by auto
moreover have 2: "S'.set (S'.dom (\<Phi> m)) \<subseteq> S'.set (S'.cod (\<Phi> m))"
using 1 \<Phi>.preserves_dom \<Phi>.preserves_cod \<Phi>_acts_elementwise by auto
moreover have "\<Phi> m =
S'.mkArr (S'.set (S'.dom (\<Phi> m))) (S'.set (S'.cod (\<Phi> m))) (\<lambda>x'. x')"
proof -
have "\<Phi> m = S'.mkArr (S'.set (\<Phi>o (S.dom m))) (S'.set (\<Phi>o (S.cod m))) (\<Phi>a m)"
using 1 \<Phi>_def by simp
also have "... = S'.mkArr (S'.set (S'.dom (\<Phi> m))) (S'.set (S'.cod (\<Phi> m))) (\<Phi>a m)"
using 1 \<Phi>_ide by auto
finally have 3: "\<Phi> m =
S'.mkArr (S'.set (S'.dom (\<Phi> m))) (S'.set (S'.cod (\<Phi> m))) (\<Phi>a m)"
by auto
also have "... = S'.mkArr (S'.set (S'.dom (\<Phi> m))) (S'.set (S'.cod (\<Phi> m))) (\<lambda>x'. x')"
proof -
have 4: "S.Fun m = restrict (\<lambda>x. x) (S.set (S.dom m))"
using assms S.incl_def by (metis (full_types) S.Fun_mkArr)
hence "\<Phi>a m = restrict (\<lambda>x'. x') (\<phi> ` (S.set (S.dom m)))"
proof -
have 5: "\<And>x'. x' \<in> \<phi> ` S.set (S.dom m) \<Longrightarrow> \<phi> (\<psi> x') = x'"
by (meson 1 S.ide_dom S.setp_imp_subset_Univ S.setp_set' f_inv_into_f
image_mono subset_eq)
have "\<Phi>a m = restrict (\<lambda>x'. \<phi> (S.Fun m (\<psi> x'))) (\<phi> ` S.set (S.dom m))"
using \<Phi>a_def by simp
also have "... = restrict (\<lambda>x'. x') (\<phi> ` S.set (S.dom m))"
proof -
have "\<And>x. x \<in> \<phi> ` (S.set (S.dom m)) \<Longrightarrow> \<phi> (S.Fun m (\<psi> x)) = x"
proof -
fix x
assume x: "x \<in> \<phi> ` (S.set (S.dom m))"
hence "\<psi> x \<in> S.set (S.dom m)"
using 1 S.ide_dom S.setp_imp_subset_Univ S.setp_set_ide \<psi>_img_\<phi>_img image_eqI
by metis
thus "\<phi> (S.Fun m (\<psi> x)) = x" using 1 4 5 x by simp
qed
thus ?thesis by auto
qed
finally show ?thesis by auto
qed
hence "\<Phi>a m = restrict (\<lambda>x'. x') (S'.set (S'.dom (\<Phi> m)))"
using 1 set_dom_\<Phi> by auto
thus ?thesis
using 2 3 \<open>S'.arr (\<Phi> m)\<close> S'.mkArr_restrict_eq S'.ide_cod S'.ide_dom S'.incl_def
by (metis S'.arr_mkArr image_restrict_eq image_subset_iff_funcset)
qed
finally show ?thesis by auto
qed
ultimately show ?thesis using S'.incl_def by blast
qed
lemma \<psi>_respects_sets:
assumes "A' \<subseteq> S'.Univ"
shows "setp (\<psi> ` A') \<longleftrightarrow> setp' A'"
using assms \<phi>_respects_setp \<phi>_img_\<psi>_img bij_\<phi>
by (metis \<psi>_img_\<phi>_img bij_betw_def image_mono order_refl)
text\<open>
Interchange the role of @{term \<phi>} and @{term \<psi>} to obtain a functor \<open>\<Psi>\<close>
from @{term[source=true] S'} to @{term[source=true] S}.
\<close>
interpretation INV: two_set_categories_bij_betw_Univ S' setp' S setp \<psi>
using \<psi>_respects_sets bij_\<phi> bij_betw_inv_into
by unfold_locales auto
abbreviation \<Psi>o
where "\<Psi>o \<equiv> INV.\<Phi>o"
abbreviation \<Psi>a
where "\<Psi>a \<equiv> INV.\<Phi>a"
abbreviation \<Psi>
where "\<Psi> \<equiv> INV.\<Phi>"
interpretation \<Psi>: "functor" S' S \<Psi>
using INV.\<Phi>_is_functor by auto
text\<open>
The functors @{term \<Phi>} and @{term \<Psi>} are inverses.
\<close>
lemma Fun_\<Psi>:
assumes "S'.arr f'" and "x' \<in> S'.set (S'.dom f')"
shows "S.Fun (\<Psi> f') (\<psi> x') = \<Psi>a f' (\<psi> x')"
using assms INV.Fun_\<Phi> by blast
lemma \<Psi>o_\<Phi>o:
assumes "S.ide a"
shows "\<Psi>o (\<Phi>o a) = a"
using assms \<Phi>o_def INV.\<Phi>o_def \<psi>_img_\<phi>_img \<Phi>o_preserves_ide set_\<Phi>o S.mkIde_set
by (simp add: S.setp_imp_subset_Univ)
lemma \<Phi>\<Psi>:
assumes "S.arr f"
shows "\<Psi> (\<Phi> f) = f"
- proof (intro S.arr_eqI)
+ proof (intro S.arr_eqI\<^sub>S\<^sub>C)
show par: "S.par (\<Psi> (\<Phi> f)) f"
using assms \<Phi>o_preserves_ide \<Psi>o_\<Phi>o by auto
show "S.Fun (\<Psi> (\<Phi> f)) = S.Fun f"
proof -
have "S.arr (\<Psi> (\<Phi> f))" using assms by auto
moreover have "\<Psi> (\<Phi> f) = S.mkArr (S.set (S.dom f)) (S.set (S.cod f)) (\<Psi>a (\<Phi> f))"
using assms INV.\<Phi>_def \<Phi>_in_hom \<Psi>o_\<Phi>o by auto
moreover have "\<Psi>a (\<Phi> f) = (\<lambda>x \<in> S.set (S.dom f). \<psi> (S'.Fun (\<Phi> f) (\<phi> x)))"
proof -
have "\<Psi>a (\<Phi> f) = (\<lambda>x \<in> \<psi> ` S'.set (S'.dom (\<Phi> f)). \<psi> (S'.Fun (\<Phi> f) (\<phi> x)))"
proof -
have "\<And>x. x \<in> \<psi> ` S'.set (S'.dom (\<Phi> f)) \<Longrightarrow> INV.\<psi> x = \<phi> x"
using assms S.ide_dom S.setp_imp_subset_Univ \<Psi>.preserves_reflects_arr par bij_\<phi>
inv_into_inv_into_eq subsetCE INV.set_dom_\<Phi>
by (metis (no_types) S.setp_set')
thus ?thesis
using INV.\<Phi>a_def by auto
qed
moreover have "\<psi> ` S'.set (S'.dom (\<Phi> f)) = S.set (S.dom f)"
using assms by (metis par \<Psi>.preserves_reflects_arr INV.set_dom_\<Phi>)
ultimately show ?thesis by auto
qed
ultimately have 1: "S.Fun (\<Psi> (\<Phi> f)) = (\<lambda>x \<in> S.set (S.dom f). \<psi> (S'.Fun (\<Phi> f) (\<phi> x)))"
using S'.Fun_mkArr by simp
show ?thesis
proof
fix x
have "x \<notin> S.set (S.dom f) \<Longrightarrow> S.Fun (\<Psi> (\<Phi> f)) x = S.Fun f x"
using 1 assms extensional_def S.Fun_mapsto S.Fun_def by auto
moreover have "x \<in> S.set (S.dom f) \<Longrightarrow> S.Fun (\<Psi> (\<Phi> f)) x = S.Fun f x"
proof -
assume x: "x \<in> S.set (S.dom f)"
have "S.Fun (\<Psi> (\<Phi> f)) x = \<psi> (\<phi> (S.Fun f (\<psi> (\<phi> x))))"
using assms x 1 Fun_\<Phi> bij_\<phi> \<Phi>a_def by auto
also have "... = S.Fun f x"
proof -
have 2: "\<And>x. x \<in> S.Univ \<Longrightarrow> \<psi> (\<phi> x) = x"
using bij_\<phi> bij_betw_inv_into_left by fast
have "S.Fun f (\<psi> (\<phi> x)) = S.Fun f x"
using assms x 2 S.ide_dom S.setp_imp_subset_Univ
by (metis S.setp_set' subsetD)
moreover have "S.Fun f x \<in> S.Univ"
using x assms S.Fun_mapsto S.setp_imp_subset_Univ S.setp_set' S.ide_cod
by blast
ultimately show ?thesis using 2 by auto
qed
finally show ?thesis by auto
qed
ultimately show "S.Fun (\<Psi> (\<Phi> f)) x = S.Fun f x" by auto
qed
qed
qed
lemma \<Phi>o_\<Psi>o:
assumes "S'.ide a'"
shows "\<Phi>o (\<Psi>o a') = a'"
using assms \<Phi>o_def INV.\<Phi>o_def \<phi>_img_\<psi>_img INV.\<Phi>o_preserves_ide \<psi>_\<phi> INV.set_\<Phi>o
S'.mkIde_set S'.setp_imp_subset_Univ
by force
lemma \<Psi>\<Phi>:
assumes "S'.arr f'"
shows "\<Phi> (\<Psi> f') = f'"
- proof (intro S'.arr_eqI)
+ proof (intro S'.arr_eqI\<^sub>S\<^sub>C)
show par: "S'.par (\<Phi> (\<Psi> f')) f'"
using assms \<Phi>.preserves_ide \<Psi>.preserves_ide \<Phi>_ide INV.\<Phi>_ide \<Phi>o_\<Psi>o by auto
show "S'.Fun (\<Phi> (\<Psi> f')) = S'.Fun f'"
proof -
have "S'.arr (\<Phi> (\<Psi> f'))" using assms by blast
moreover have "\<Phi> (\<Psi> f') =
S'.mkArr (S'.set (S'.dom f')) (S'.set (S'.cod f')) (\<Phi>a (\<Psi> f'))"
using assms \<Phi>_def INV.\<Phi>_in_hom \<Phi>o_\<Psi>o by simp
moreover have "\<Phi>a (\<Psi> f') = (\<lambda>x' \<in> S'.set (S'.dom f'). \<phi> (S.Fun (\<Psi> f') (\<psi> x')))"
unfolding \<Phi>a_def
using assms par \<Psi>.preserves_arr set_dom_\<Phi> by metis
ultimately have 1: "S'.Fun (\<Phi> (\<Psi> f')) =
(\<lambda>x' \<in> S'.set (S'.dom f'). \<phi> (S.Fun (\<Psi> f') (\<psi> x')))"
using S'.Fun_mkArr by simp
show ?thesis
proof
fix x'
have "x' \<notin> S'.set (S'.dom f') \<Longrightarrow> S'.Fun (\<Phi> (\<Psi> f')) x' = S'.Fun f' x'"
using 1 assms S'.Fun_mapsto extensional_def by (simp add: S'.Fun_def)
moreover have "x' \<in> S'.set (S'.dom f') \<Longrightarrow> S'.Fun (\<Phi> (\<Psi> f')) x' = S'.Fun f' x'"
proof -
assume x': "x' \<in> S'.set (S'.dom f')"
have "S'.Fun (\<Phi> (\<Psi> f')) x' = \<phi> (S.Fun (\<Psi> f') (\<psi> x'))"
using x' 1 by auto
also have "... = \<phi> (\<Psi>a f' (\<psi> x'))"
using Fun_\<Psi> x' assms S'.setp_imp_subset_Univ bij_\<phi> by metis
also have "... = \<phi> (\<psi> (S'.Fun f' (\<phi> (\<psi> x'))))"
proof -
have "\<phi> (\<Psi>a f' (\<psi> x')) = \<phi> (\<psi> (S'.Fun f' x'))"
proof -
have "x' \<in> S'.Univ"
by (meson S'.ide_dom S'.setp_imp_subset_Univ S'.setp_set_ide assms subsetCE x')
thus ?thesis
by (simp add: INV.\<Phi>a_def INV.\<psi>_\<phi> x')
qed
also have "... = \<phi> (\<psi> (S'.Fun f' (\<phi> (\<psi> x'))))"
using assms x' \<phi>_\<psi> S'.setp_imp_subset_Univ S'.setp_set_ide S'.ide_dom
by (metis subsetCE)
finally show ?thesis by auto
qed
also have "... = S'.Fun f' x'"
proof -
have 2: "\<And>x'. x' \<in> S'.Univ \<Longrightarrow> \<phi> (\<psi> x') = x'"
using bij_\<phi> bij_betw_inv_into_right by fast
have "S'.Fun f' (\<phi> (\<psi> x')) = S'.Fun f' x'"
using assms x' 2 S'.setp_imp_subset_Univ S'.setp_set_ide S'.ide_dom
by (metis subsetCE)
moreover have "S'.Fun f' x' \<in> S'.Univ"
using x' assms S'.Fun_mapsto S'.setp_imp_subset_Univ S'.setp_set_ide S'.ide_cod
by blast
ultimately show ?thesis using 2 by auto
qed
finally show ?thesis by auto
qed
ultimately show "S'.Fun (\<Phi> (\<Psi> f')) x' = S'.Fun f' x'" by auto
qed
qed
qed
lemma inverse_functors_\<Phi>_\<Psi>:
shows "inverse_functors S S' \<Psi> \<Phi>"
proof -
interpret \<Phi>\<Psi>: composite_functor S S' S \<Phi> \<Psi> ..
have inv: "\<Psi> o \<Phi> = S.map"
using \<Phi>\<Psi> S.map_def \<Phi>\<Psi>.is_extensional by auto
interpret \<Psi>\<Phi>: composite_functor S' S S' \<Psi> \<Phi> ..
have inv': "\<Phi> o \<Psi> = S'.map"
using \<Psi>\<Phi> S'.map_def \<Psi>\<Phi>.is_extensional by auto
show ?thesis
using inv inv' by (unfold_locales, auto)
qed
lemma are_isomorphic:
shows "\<exists>\<Phi>. invertible_functor S S' \<Phi> \<and> (\<forall>m. S.incl m \<longrightarrow> S'.incl (\<Phi> m))"
proof -
interpret inverse_functors S S' \<Psi> \<Phi>
using inverse_functors_\<Phi>_\<Psi> by auto
have 1: "inverse_functors S S' \<Psi> \<Phi>" ..
interpret invertible_functor S S' \<Phi>
apply unfold_locales using 1 by auto
have "invertible_functor S S' \<Phi>" ..
thus ?thesis using \<Phi>_preserves_incl by auto
qed
end
text\<open>
The main result: @{locale set_category} is categorical, in the following (logical) sense:
If \<open>S\<close> and \<open>S'\<close> are two "set categories", and if the sets of terminal objects of \<open>S\<close> and \<open>S'\<close>
are in correspondence via a @{term setp}-preserving bijection, then \<open>S\<close> and \<open>S'\<close> are
isomorphic as categories, via a functor that preserves inclusion maps, hence also the
inclusion relation between sets.
\<close>
theorem set_category_is_categorical:
assumes "set_category S setp" and "set_category S' setp'"
and "bij_betw \<phi> (set_category_data.Univ S) (set_category_data.Univ S')"
and "\<And>A. A \<subseteq> set_category_data.Univ S \<Longrightarrow> setp' (\<phi> ` A) \<longleftrightarrow> setp A"
shows "\<exists>\<Phi>. invertible_functor S S' \<Phi> \<and>
(\<forall>m. set_category.incl S setp m \<longrightarrow> set_category.incl S' setp' (\<Phi> m))"
proof -
interpret S: set_category S setp using assms(1) by auto
interpret S': set_category S' setp' using assms(2) by auto
interpret two_set_categories_bij_betw_Univ S setp S' setp' \<phi>
apply (unfold_locales) using assms(3-4) by auto
show ?thesis using are_isomorphic by auto
qed
section "Further Properties of Set Categories"
text\<open>
In this section we further develop the consequences of the \<open>set_category\<close>
axioms, and establish characterizations of a number of standard category-theoretic
notions for a \<open>set_category\<close>.
\<close>
context set_category
begin
abbreviation Dom
where "Dom f \<equiv> set (dom f)"
abbreviation Cod
where "Cod f \<equiv> set (cod f)"
subsection "Initial Object"
text\<open>
The object corresponding to the empty set is an initial object.
\<close>
definition empty
where "empty = mkIde {}"
lemma initial_empty:
shows "initial empty"
proof
show 0: "ide empty"
using empty_def by (simp add: setp_empty)
show "\<And>b. ide b \<Longrightarrow> \<exists>!f. \<guillemotleft>f : empty \<rightarrow> b\<guillemotright>"
proof -
fix b
assume b: "ide b"
show "\<exists>!f. \<guillemotleft>f : empty \<rightarrow> b\<guillemotright>"
proof
show 1: "\<guillemotleft>mkArr {} (set b) (\<lambda>x. x) : empty \<rightarrow> b\<guillemotright>"
using 0 b empty_def mkArr_in_hom mkIde_set setp_imp_subset_Univ arr_mkIde
by (metis Pi_I empty_iff ide_def mkIde_def)
show "\<And>f. \<guillemotleft>f : empty \<rightarrow> b\<guillemotright> \<Longrightarrow> f = mkArr {} (set b) (\<lambda>x. x)"
by (metis "1" arr_mkArr empty_iff in_homE empty_def mkArr_Fun mkArr_eqI set_mkIde)
qed
qed
qed
subsection "Identity Arrows"
text\<open>
Identity arrows correspond to restrictions of the identity function.
\<close>
- lemma ide_char:
+ lemma ide_char\<^sub>S\<^sub>C:
assumes "arr f"
shows "ide f \<longleftrightarrow> Dom f = Cod f \<and> Fun f = (\<lambda>x \<in> Dom f. x)"
using assms mkIde_as_mkArr mkArr_Fun Fun_ide in_homE ide_cod mkArr_Fun mkIde_set
by (metis ide_char)
lemma ideI:
assumes "arr f" and "Dom f = Cod f" and "\<And>x. x \<in> Dom f \<Longrightarrow> Fun f x = x"
shows "ide f"
proof -
have "Fun f = (\<lambda>x \<in> Dom f. x)"
using assms Fun_def by auto
- thus ?thesis using assms ide_char by blast
+ thus ?thesis using assms ide_char\<^sub>S\<^sub>C by blast
qed
subsection "Inclusions"
lemma ide_implies_incl:
assumes "ide a"
shows "incl a"
by (simp add: assms incl_def)
definition incl_in :: "'s \<Rightarrow> 's \<Rightarrow> bool"
where "incl_in a b = (ide a \<and> ide b \<and> set a \<subseteq> set b)"
abbreviation incl_of
where "incl_of a b \<equiv> mkArr (set a) (set b) (\<lambda>x. x)"
lemma elem_set_implies_set_eq_singleton:
assumes "a \<in> set b"
shows "set a = {a}"
proof -
have "ide b" using assms set_def by auto
thus ?thesis using assms setp_imp_subset_Univ terminal_char2
by (metis setp_set' insert_subset mem_Collect_eq mk_disjoint_insert)
qed
lemma elem_set_implies_incl_in:
assumes "a \<in> set b"
shows "incl_in a b"
proof -
have b: "ide b" using assms set_def by auto
hence "setp (set b)" by simp
hence "a \<in> Univ \<and> set a \<subseteq> set b"
using setp_imp_subset_Univ assms elem_set_implies_set_eq_singleton by auto
hence "ide a \<and> set a \<subseteq> set b"
using b terminal_char1 by simp
thus ?thesis using b incl_in_def by simp
qed
lemma incl_incl_of [simp]:
assumes "incl_in a b"
shows "incl (incl_of a b)"
and "\<guillemotleft>incl_of a b : a \<rightarrow> b\<guillemotright>"
proof -
show "\<guillemotleft>incl_of a b : a \<rightarrow> b\<guillemotright>"
using assms incl_in_def mkArr_in_hom
by (metis image_ident image_subset_iff_funcset mkIde_set setp_set_ide)
thus "incl (incl_of a b)"
using assms incl_def incl_in_def by fastforce
qed
text\<open>
There is at most one inclusion between any pair of objects.
\<close>
lemma incls_coherent:
assumes "par f f'" and "incl f" and "incl f'"
shows "f = f'"
using assms incl_def fun_complete by auto
text\<open>
The set of inclusions is closed under composition.
\<close>
lemma incl_comp [simp]:
assumes "incl f" and "incl g" and "cod f = dom g"
shows "incl (g \<cdot> f)"
proof -
have 1: "seq g f" using assms incl_def by auto
moreover have 2: "Dom (g \<cdot> f) \<subseteq> Cod (g \<cdot> f)"
using assms 1 incl_def by auto
moreover have "g \<cdot> f = mkArr (Dom f) (Cod g) (restrict (\<lambda>x. x) (Dom f))"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 3: "arr (mkArr (Dom f) (Cod g) (\<lambda>x\<in>Dom f. x))"
by (metis 1 2 cod_comp dom_comp ide_cod ide_dom incl_def incl_in_def
incl_incl_of(1) mkArr_restrict_eq)
show 4: "par (g \<cdot> f) (mkArr (Dom f) (Cod g) (\<lambda>x\<in>Dom f. x))"
using assms 1 3 mkIde_set by auto
show "Fun (g \<cdot> f) = Fun (mkArr (Dom f) (Cod g) (\<lambda>x\<in>Dom f. x))"
using assms 3 4 Fun_comp Fun_mkArr
by (metis comp_cod_arr dom_cod ide_cod ide_implies_incl incl_def mkArr_restrict_eq')
qed
ultimately show ?thesis using incl_def by force
qed
subsection "Image Factorization"
text\<open>
The image of an arrow is the object that corresponds to the set-theoretic
image of the domain set under the function induced by the arrow.
\<close>
abbreviation Img
where "Img f \<equiv> Fun f ` Dom f"
definition img
where "img f = mkIde (Img f)"
lemma ide_img [simp]:
assumes "arr f"
shows "ide (img f)"
proof -
have "Fun f ` Dom f \<subseteq> Cod f" using assms Fun_mapsto by blast
moreover have "setp (Cod f)" using assms by simp
ultimately show ?thesis using img_def setp_respects_subset by auto
qed
lemma set_img [simp]:
assumes "arr f"
shows "set (img f) = Img f"
proof -
have 1: "Img f \<subseteq> Cod f \<and> setp (set (cod f))"
using assms Fun_mapsto by auto
hence "Fun f ` set (dom f) \<subseteq> Univ"
using setp_imp_subset_Univ by blast
thus ?thesis
using assms 1 img_def set_mkIde setp_respects_subset by auto
qed
lemma img_point_in_Univ:
assumes "\<guillemotleft>x : unity \<rightarrow> a\<guillemotright>"
shows "img x \<in> Univ"
proof -
have "set (img x) = {Fun x unity}"
- using assms terminal_char2 terminal_unity by auto
+ using assms terminal_char2 terminal_unity\<^sub>S\<^sub>C by auto
thus "img x \<in> Univ" using assms terminal_char1 by auto
qed
lemma incl_in_img_cod:
assumes "arr f"
shows "incl_in (img f) (cod f)"
proof (unfold img_def)
have 1: "Img f \<subseteq> Cod f \<and> setp (Cod f)"
using assms Fun_mapsto by auto
hence 2: "ide (mkIde (Img f))"
using setp_respects_subset by auto
moreover have "ide (cod f)" using assms by auto
moreover have "set (mkIde (Img f)) \<subseteq> Cod f"
using 1 2 using setp_respects_subset by force
ultimately show "incl_in (mkIde (Img f)) (cod f)"
using assms incl_in_def ide_cod by blast
qed
lemma img_point_elem_set:
assumes "\<guillemotleft>x : unity \<rightarrow> a\<guillemotright>"
shows "img x \<in> set a"
by (metis assms img_point_in_Univ in_homE incl_in_img_cod insert_subset
mem_Collect_eq incl_in_def terminal_char2)
text\<open>
The corestriction of an arrow @{term f} is the arrow
@{term "corestr f \<in> hom (dom f) (img f)"} that induces the same function
on the universe as @{term f}.
\<close>
definition corestr
where "corestr f = mkArr (Dom f) (Img f) (Fun f)"
lemma corestr_in_hom:
assumes "arr f"
shows "\<guillemotleft>corestr f : dom f \<rightarrow> img f\<guillemotright>"
by (metis assms corestr_def equalityD2 ide_dom ide_img image_subset_iff_funcset
mkIde_set set_img mkArr_in_hom setp_set_ide)
text\<open>
Every arrow factors as a corestriction followed by an inclusion.
\<close>
lemma img_fact:
assumes "arr f"
shows "S (incl_of (img f) (cod f)) (corestr f) = f"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 1: "\<guillemotleft>corestr f : dom f \<rightarrow> img f\<guillemotright>"
using assms corestr_in_hom by blast
moreover have 2: "\<guillemotleft>incl_of (img f) (cod f) : img f \<rightarrow> cod f\<guillemotright>"
using assms incl_in_img_cod incl_incl_of by fast
ultimately show P: "par (incl_of (img f) (cod f) \<cdot> corestr f) f"
using assms in_homE by blast
show "Fun (incl_of (img f) (cod f) \<cdot> corestr f) = Fun f"
by (metis (no_types, lifting) 1 2 Fun_comp Fun_ide Fun_mkArr P comp_cod_arr
corestr_def ide_img in_homE mkArr_Fun)
qed
lemma Fun_corestr:
assumes "arr f"
shows "Fun (corestr f) = Fun f"
by (metis Fun_mkArr arrI assms corestr_def corestr_in_hom mkArr_Fun)
subsection "Points and Terminal Objects"
text\<open>
To each element @{term t} of @{term "set a"} is associated a point
@{term "mkPoint a t \<in> hom unity a"}. The function induced by such
a point is the constant-@{term t} function on the set @{term "{unity}"}.
\<close>
definition mkPoint
where "mkPoint a t \<equiv> mkArr {unity} (set a) (\<lambda>_. t)"
lemma mkPoint_in_hom:
assumes "ide a" and "t \<in> set a"
shows "\<guillemotleft>mkPoint a t : unity \<rightarrow> a\<guillemotright>"
using assms mkArr_in_hom
- by (metis Pi_I mkIde_set setp_set_ide terminal_char2 terminal_unity mkPoint_def)
+ by (metis Pi_I mkIde_set setp_set_ide terminal_char2 terminal_unity\<^sub>S\<^sub>C mkPoint_def)
lemma Fun_mkPoint:
assumes "ide a" and "t \<in> set a"
shows "Fun (mkPoint a t) = (\<lambda>_ \<in> {unity}. t)"
- using assms mkPoint_def terminal_unity mkPoint_in_hom by fastforce
+ using assms mkPoint_def terminal_unity\<^sub>S\<^sub>C mkPoint_in_hom by fastforce
text\<open>
For each object @{term a} the function @{term "mkPoint a"} has as its inverse
the restriction of the function @{term img} to @{term "hom unity a"}
\<close>
lemma mkPoint_img:
shows "img \<in> hom unity a \<rightarrow> set a"
and "\<And>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<Longrightarrow> mkPoint a (img x) = x"
proof -
show "img \<in> hom unity a \<rightarrow> set a"
using img_point_elem_set by simp
show "\<And>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<Longrightarrow> mkPoint a (img x) = x"
proof -
fix x
assume x: "\<guillemotleft>x : unity \<rightarrow> a\<guillemotright>"
show "mkPoint a (img x) = x"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 0: "img x \<in> set a"
using x img_point_elem_set by metis
hence 1: "mkPoint a (img x) \<in> hom unity a"
using x mkPoint_in_hom by force
thus 2: "par (mkPoint a (img x)) x"
using x by fastforce
have "Fun (mkPoint a (img x)) = (\<lambda>_ \<in> {unity}. img x)"
using 1 mkPoint_def by auto
also have "... = Fun x"
by (metis 0 Fun_corestr calculation elem_set_implies_set_eq_singleton
ide_cod ide_unity in_homE mem_Collect_eq Fun_mkPoint corestr_in_hom
img_point_in_Univ mkPoint_in_hom singletonI terminalE x)
finally show "Fun (mkPoint a (img x)) = Fun x" by auto
qed
qed
qed
lemma img_mkPoint:
assumes "ide a"
shows "mkPoint a \<in> set a \<rightarrow> hom unity a"
and "\<And>t. t \<in> set a \<Longrightarrow> img (mkPoint a t) = t"
proof -
show "mkPoint a \<in> set a \<rightarrow> hom unity a"
using assms(1) mkPoint_in_hom by simp
show "\<And>t. t \<in> set a \<Longrightarrow> img (mkPoint a t) = t"
proof -
fix t
assume t: "t \<in> set a"
show "img (mkPoint a t) = t"
proof -
have 1: "arr (mkPoint a t)"
using assms t mkPoint_in_hom by auto
have "Fun (mkPoint a t) ` {unity} = {t}"
using 1 mkPoint_def by simp
thus ?thesis
by (metis in_homE img_def mkIde_set mkPoint_in_hom elem_set_implies_incl_in
- elem_set_implies_set_eq_singleton incl_in_def t terminal_char2 terminal_unity)
+ elem_set_implies_set_eq_singleton incl_in_def t terminal_char2 terminal_unity\<^sub>S\<^sub>C)
qed
qed
qed
text\<open>
For each object @{term a} the elements of @{term "hom unity a"} are therefore in
bijective correspondence with @{term "set a"}.
\<close>
lemma bij_betw_points_and_set:
assumes "ide a"
shows "bij_betw img (hom unity a) (set a)"
proof (intro bij_betwI)
show "img \<in> hom unity a \<rightarrow> set a"
using assms mkPoint_img by auto
show "mkPoint a \<in> set a \<rightarrow> hom unity a"
using assms img_mkPoint by auto
show "\<And>x. x \<in> hom unity a \<Longrightarrow> mkPoint a (img x) = x"
using assms mkPoint_img by auto
show "\<And>t. t \<in> set a \<Longrightarrow> img (mkPoint a t) = t"
using assms img_mkPoint by auto
qed
lemma setp_img_points:
assumes "ide a"
shows "setp (img ` hom unity a)"
using assms
by (metis image_subset_iff_funcset mkPoint_img(1) setp_respects_subset setp_set_ide)
text\<open>
The function on the universe induced by an arrow @{term f} agrees, under the bijection
between @{term "hom unity (dom f)"} and @{term "Dom f"}, with the action of @{term f}
by composition on @{term "hom unity (dom f)"}.
\<close>
lemma Fun_point:
assumes "\<guillemotleft>x : unity \<rightarrow> a\<guillemotright>"
shows "Fun x = (\<lambda>_ \<in> {unity}. img x)"
using assms mkPoint_img img_mkPoint Fun_mkPoint [of a "img x"] img_point_elem_set
by auto
lemma comp_arr_mkPoint:
assumes "arr f" and "t \<in> Dom f"
shows "f \<cdot> mkPoint (dom f) t = mkPoint (cod f) (Fun f t)"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
have 0: "seq f (mkPoint (dom f) t)"
using assms mkPoint_in_hom [of "dom f" t] by auto
have 1: "\<guillemotleft>f \<cdot> mkPoint (dom f) t : unity \<rightarrow> cod f\<guillemotright>"
using assms mkPoint_in_hom [of "dom f" t] by auto
show "par (f \<cdot> mkPoint (dom f) t) (mkPoint (cod f) (Fun f t))"
proof -
have "\<guillemotleft>mkPoint (cod f) (Fun f t) : unity \<rightarrow> cod f\<guillemotright>"
using assms Fun_mapsto mkPoint_in_hom [of "cod f" "Fun f t"] by auto
thus ?thesis using 1 by fastforce
qed
show "Fun (f \<cdot> mkPoint (dom f) t) = Fun (mkPoint (cod f) (Fun f t))"
proof -
have "Fun (f \<cdot> mkPoint (dom f) t) = restrict (Fun f o Fun (mkPoint (dom f) t)) {unity}"
- using assms 0 1 Fun_comp terminal_char2 terminal_unity by auto
+ using assms 0 1 Fun_comp terminal_char2 terminal_unity\<^sub>S\<^sub>C by auto
also have "... = (\<lambda>_ \<in> {unity}. Fun f t)"
using assms Fun_mkPoint by auto
also have "... = Fun (mkPoint (cod f) (Fun f t))"
using assms Fun_mkPoint [of "cod f" "Fun f t"] Fun_mapsto by fastforce
finally show ?thesis by auto
qed
qed
- lemma comp_arr_point:
+ lemma comp_arr_point\<^sub>S\<^sub>S\<^sub>C:
assumes "arr f" and "\<guillemotleft>x : unity \<rightarrow> dom f\<guillemotright>"
shows "f \<cdot> x = mkPoint (cod f) (Fun f (img x))"
by (metis assms comp_arr_mkPoint img_point_elem_set mkPoint_img(2))
text\<open>
This agreement allows us to express @{term "Fun f"} in terms of composition.
\<close>
lemma Fun_in_terms_of_comp:
assumes "arr f"
shows "Fun f = restrict (img o S f o mkPoint (dom f)) (Dom f)"
proof
fix t
have "t \<notin> Dom f \<Longrightarrow> Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t"
using assms by (simp add: Fun_def)
moreover have "t \<in> Dom f \<Longrightarrow>
Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t"
proof -
assume t: "t \<in> Dom f"
have 1: "f \<cdot> mkPoint (dom f) t = mkPoint (cod f) (Fun f t)"
using assms t comp_arr_mkPoint by simp
hence "img (f \<cdot> mkPoint (dom f) t) = img (mkPoint (cod f) (Fun f t))" by simp
thus ?thesis
proof -
have "Fun f t \<in> Cod f" using assms t Fun_mapsto by auto
thus ?thesis using assms t 1 img_mkPoint by auto
qed
qed
ultimately show "Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" by auto
qed
text\<open>
We therefore obtain a rule for proving parallel arrows equal by showing
that they have the same action by composition on points.
\<close>
(*
* TODO: Find out why attempts to use this as the main rule for a proof loop
* unless the specific instance is given.
*)
- lemma arr_eqI':
+ lemma arr_eqI'\<^sub>S\<^sub>C:
assumes "par f f'" and "\<And>x. \<guillemotleft>x : unity \<rightarrow> dom f\<guillemotright> \<Longrightarrow> f \<cdot> x = f' \<cdot> x"
shows "f = f'"
- using assms Fun_in_terms_of_comp mkPoint_in_hom by (intro arr_eqI, auto)
+ using assms Fun_in_terms_of_comp mkPoint_in_hom by (intro arr_eqI\<^sub>S\<^sub>C, auto)
text\<open>
An arrow can therefore be specified by giving its action by composition on points.
In many situations, this is more natural than specifying it as a function on the universe.
\<close>
definition mkArr'
where "mkArr' a b F = mkArr (set a) (set b) (img o F o mkPoint a)"
lemma mkArr'_in_hom:
assumes "ide a" and "ide b" and "F \<in> hom unity a \<rightarrow> hom unity b"
shows "\<guillemotleft>mkArr' a b F : a \<rightarrow> b\<guillemotright>"
proof -
have "img o F o mkPoint a \<in> set a \<rightarrow> set b"
using assms(1,3) img_mkPoint(1) mkPoint_img(1) by fastforce
thus ?thesis
using assms mkArr'_def mkArr_in_hom [of "set a" "set b"] mkIde_set by simp
qed
lemma comp_point_mkArr':
assumes "ide a" and "ide b" and "F \<in> hom unity a \<rightarrow> hom unity b"
shows "\<And>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<Longrightarrow> mkArr' a b F \<cdot> x = F x"
proof -
fix x
assume x: "\<guillemotleft>x : unity \<rightarrow> a\<guillemotright>"
have "Fun (mkArr' a b F) (img x) = img (F x)"
unfolding mkArr'_def
using assms x Fun_mkArr img_point_elem_set mkPoint_img mkPoint_in_hom
by (simp add: Pi_iff)
hence "mkArr' a b F \<cdot> x = mkPoint b (img (F x))"
- using assms x mkArr'_in_hom [of a b F] comp_arr_point by auto
+ using assms x mkArr'_in_hom [of a b F] comp_arr_point\<^sub>S\<^sub>S\<^sub>C by auto
thus "mkArr' a b F \<cdot> x = F x"
using assms x mkPoint_img(2) by auto
qed
text\<open>
A third characterization of terminal objects is as those objects whose set of
points is a singleton.
\<close>
lemma terminal_char3:
assumes "\<exists>!x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright>"
shows "terminal a"
proof -
have a: "ide a"
using assms ide_cod mem_Collect_eq by blast
hence "bij_betw img (hom unity a) (set a)"
using assms bij_betw_points_and_set by auto
hence "img ` (hom unity a) = set a"
by (simp add: bij_betw_def)
moreover have "hom unity a = {THE x. x \<in> hom unity a}"
using assms theI' [of "\<lambda>x. x \<in> hom unity a"] by auto
ultimately have "set a = {img (THE x. x \<in> hom unity a)}"
by (metis image_empty image_insert)
thus ?thesis using a terminal_char1 by simp
qed
text\<open>
The following is an alternative formulation of functional completeness, which says that
any function on points uniquely determines an arrow.
\<close>
lemma fun_complete':
assumes "ide a" and "ide b" and "F \<in> hom unity a \<rightarrow> hom unity b"
shows "\<exists>!f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<longrightarrow> f \<cdot> x = F x)"
proof
have 1: "\<guillemotleft>mkArr' a b F : a \<rightarrow> b\<guillemotright>" using assms mkArr'_in_hom by auto
moreover have 2: "\<And>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<Longrightarrow> mkArr' a b F \<cdot> x = F x"
using assms comp_point_mkArr' by auto
ultimately show "\<guillemotleft>mkArr' a b F : a \<rightarrow> b\<guillemotright> \<and>
(\<forall>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<longrightarrow> mkArr' a b F \<cdot> x = F x)" by blast
fix f
assume f: "\<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> (\<forall>x. \<guillemotleft>x : unity \<rightarrow> a\<guillemotright> \<longrightarrow> f \<cdot> x = F x)"
show "f = mkArr' a b F"
- using f 1 2 by (intro arr_eqI' [of f "mkArr' a b F"], fastforce, auto)
+ using f 1 2 by (intro arr_eqI'\<^sub>S\<^sub>C [of f "mkArr' a b F"], fastforce, auto)
qed
subsection "The `Determines Same Function' Relation on Arrows"
text\<open>
An important part of understanding the structure of a category of sets and functions
is to characterize when it is that two arrows ``determine the same function''.
The following result provides one answer to this: two arrows with a common domain
determine the same function if and only if they can be rendered equal by composing with
a cospan of inclusions.
\<close>
lemma eq_Fun_iff_incl_joinable:
assumes "span f f'"
shows "Fun f = Fun f' \<longleftrightarrow>
(\<exists>m m'. incl m \<and> incl m' \<and> seq m f \<and> seq m' f' \<and> m \<cdot> f = m' \<cdot> f')"
proof
assume ff': "Fun f = Fun f'"
let ?b = "mkIde (Cod f \<union> Cod f')"
let ?m = "incl_of (cod f) ?b"
let ?m' = "incl_of (cod f') ?b"
have incl_m: "incl ?m"
using assms incl_incl_of [of "cod f" ?b] incl_in_def setp_respects_union by simp
have incl_m': "incl ?m'"
using assms incl_incl_of [of "cod f'" ?b] incl_in_def setp_respects_union by simp
have m: "?m = mkArr (Cod f) (Cod f \<union> Cod f') (\<lambda>x. x)"
using setp_respects_union by (simp add: assms)
have m': "?m' = mkArr (Cod f') (Cod f \<union> Cod f') (\<lambda>x. x)"
using setp_respects_union by (simp add: assms)
have seq: "seq ?m f \<and> seq ?m' f'"
using assms m m' using setp_respects_union by simp
have "?m \<cdot> f = ?m' \<cdot> f'"
by (metis assms comp_mkArr ff' incl_def incl_m incl_m' mkArr_Fun)
hence "incl ?m \<and> incl ?m' \<and> seq ?m f \<and> seq ?m' f' \<and> ?m \<cdot> f = ?m' \<cdot> f'"
using seq \<open>incl ?m\<close> \<open>incl ?m'\<close> by simp
thus "\<exists>m m'. incl m \<and> incl m' \<and> seq m f \<and> seq m' f' \<and> m \<cdot> f = m' \<cdot> f'" by auto
next
assume ff': "\<exists>m m'. incl m \<and> incl m' \<and> seq m f \<and> seq m' f' \<and> m \<cdot> f = m' \<cdot> f'"
show "Fun f = Fun f'"
using ff'
by (metis Fun_comp Fun_ide comp_cod_arr ide_cod seqE Fun_incl)
qed
text\<open>
Another answer to the same question: two arrows with a common domain determine the
same function if and only if their corestrictions are equal.
\<close>
lemma eq_Fun_iff_eq_corestr:
assumes "span f f'"
shows "Fun f = Fun f' \<longleftrightarrow> corestr f = corestr f'"
using assms corestr_def Fun_corestr by metis
subsection "Retractions, Sections, and Isomorphisms"
text\<open>
An arrow is a retraction if and only if its image coincides with its codomain.
\<close>
lemma retraction_if_Img_eq_Cod:
assumes "arr g" and "Img g = Cod g"
shows "retraction g"
and "ide (g \<cdot> mkArr (Cod g) (Dom g) (inv_into (Dom g) (Fun g)))"
proof -
let ?F = "inv_into (Dom g) (Fun g)"
let ?f = "mkArr (Cod g) (Dom g) ?F"
have f: "arr ?f"
by (simp add: assms inv_into_into)
show "ide (g \<cdot> ?f)"
proof -
have "g = mkArr (Dom g) (Cod g) (Fun g)" using assms mkArr_Fun by auto
hence "g \<cdot> ?f = mkArr (Cod g) (Cod g) (Fun g o ?F)"
using assms(1) f comp_mkArr by metis
moreover have "mkArr (Cod g) (Cod g) (\<lambda>y. y) = ..."
proof (intro mkArr_eqI')
show "arr (mkArr (Cod g) (Cod g) (\<lambda>y. y))"
using assms arr_cod_iff_arr by auto
show "\<And>y. y \<in> Cod g \<Longrightarrow> y = (Fun g o ?F) y"
using assms by (simp add: f_inv_into_f)
qed
ultimately show ?thesis
using assms f mkIde_as_mkArr by auto
qed
thus "retraction g" by auto
qed
lemma retraction_char:
shows "retraction g \<longleftrightarrow> arr g \<and> Img g = Cod g"
proof
assume G: "retraction g"
show "arr g \<and> Img g = Cod g"
proof
show "arr g" using G by blast
show "Img g = Cod g"
proof -
from G obtain f where f: "ide (g \<cdot> f)" by blast
have "Fun g ` Fun f ` Cod g = Cod g"
proof -
have "restrict (Fun g o Fun f) (Cod g) = restrict (\<lambda>x. x) (Cod g)"
using f Fun_comp Fun_ide ide_compE by metis
thus ?thesis
by (metis image_comp image_ident image_restrict_eq)
qed
moreover have "Fun f ` Cod g \<subseteq> Dom g"
using f Fun_mapsto arr_mkArr mkArr_Fun funcset_image
by (metis seqE ide_compE ide_compE)
moreover have "Img g \<subseteq> Cod g"
using f Fun_mapsto by blast
ultimately show ?thesis by blast
qed
qed
next
assume "arr g \<and> Img g = Cod g"
thus "retraction g" using retraction_if_Img_eq_Cod by blast
qed
text\<open>
Every corestriction is a retraction.
\<close>
lemma retraction_corestr:
assumes "arr f"
shows "retraction (corestr f)"
using assms retraction_char Fun_corestr corestr_in_hom in_homE set_img
by metis
text\<open>
An arrow is a section if and only if it induces an injective function on its
domain, except in the special case that it has an empty domain set and a
nonempty codomain set.
\<close>
lemma section_if_inj:
assumes "arr f" and "inj_on (Fun f) (Dom f)" and "Dom f = {} \<longrightarrow> Cod f = {}"
shows "section f"
and "ide (mkArr (Cod f) (Dom f)
(\<lambda>y. if y \<in> Img f then SOME x. x \<in> Dom f \<and> Fun f x = y
else SOME x. x \<in> Dom f)
\<cdot> f)"
proof -
let ?P= "\<lambda>y. \<lambda>x. x \<in> Dom f \<and> Fun f x = y"
let ?G = "\<lambda>y. if y \<in> Img f then SOME x. ?P y x else SOME x. x \<in> Dom f"
let ?g = "mkArr (Cod f) (Dom f) ?G"
have g: "arr ?g"
proof -
have 1: "setp (Cod f)" using assms by simp
have 2: "setp (Dom f)" using assms by simp
have 3: "?G \<in> Cod f \<rightarrow> Dom f"
proof
fix y
assume Y: "y \<in> Cod f"
show "?G y \<in> Dom f"
proof (cases "y \<in> Img f")
assume "y \<in> Img f"
hence "(\<exists>x. ?P y x) \<and> ?G y = (SOME x. ?P y x)" using Y by auto
hence "?P y (?G y)" using someI_ex [of "?P y"] by argo
thus "?G y \<in> Dom f" by auto
next
assume "y \<notin> Img f"
hence "(\<exists>x. x \<in> Dom f) \<and> ?G y = (SOME x. x \<in> Dom f)" using assms Y by auto
thus "?G y \<in> Dom f" using someI_ex [of "\<lambda>x. x \<in> Dom f"] by argo
qed
qed
show ?thesis using 1 2 3 by simp
qed
show "ide (?g \<cdot> f)"
proof -
have "f = mkArr (Dom f) (Cod f) (Fun f)" using assms mkArr_Fun by auto
hence "?g \<cdot> f = mkArr (Dom f) (Dom f) (?G o Fun f)"
using assms(1) g comp_mkArr [of "Dom f" "Cod f" "Fun f" "Dom f" ?G] by argo
moreover have "mkArr (Dom f) (Dom f) (\<lambda>x. x) = ..."
proof (intro mkArr_eqI')
show "arr (mkArr (Dom f) (Dom f) (\<lambda>x. x))"
using assms by auto
show "\<And>x. x \<in> Dom f \<Longrightarrow> x = (?G o Fun f) x"
proof -
fix x
assume x: "x \<in> Dom f"
have "Fun f x \<in> Img f" using x by blast
hence *: "(\<exists>x'. ?P (Fun f x) x') \<and> ?G (Fun f x) = (SOME x'. ?P (Fun f x) x')"
by auto
then have "?P (Fun f x) (?G (Fun f x))"
using someI_ex [of "?P (Fun f x)"] by argo
with * have "x = ?G (Fun f x)"
using assms x inj_on_def [of "Fun f" "Dom f"] by simp
thus "x = (?G o Fun f) x" by simp
qed
qed
ultimately show ?thesis
using assms mkIde_as_mkArr mkIde_set by auto
qed
thus "section f" by auto
qed
lemma section_char:
shows "section f \<longleftrightarrow> arr f \<and> (Dom f = {} \<longrightarrow> Cod f = {}) \<and> inj_on (Fun f) (Dom f)"
proof
assume f: "section f"
from f obtain g where g: "ide (g \<cdot> f)" using section_def by blast
show "arr f \<and> (Dom f = {} \<longrightarrow> Cod f = {}) \<and> inj_on (Fun f) (Dom f)"
proof -
have "arr f" using f by blast
moreover have "Dom f = {} \<longrightarrow> Cod f = {}"
proof -
have "Cod f \<noteq> {} \<longrightarrow> Dom f \<noteq> {}"
proof
assume "Cod f \<noteq> {}"
from this obtain y where "y \<in> Cod f" by blast
hence "Fun g y \<in> Dom f"
using g Fun_mapsto
by (metis seqE ide_compE image_eqI retractionI retraction_char)
thus "Dom f \<noteq> {}" by blast
qed
thus ?thesis by auto
qed
moreover have "inj_on (Fun f) (Dom f)"
by (metis Fun_comp Fun_ide g ide_compE inj_on_id2 inj_on_imageI2 inj_on_restrict_eq)
ultimately show ?thesis by auto
qed
next
assume F: "arr f \<and> (Dom f = {} \<longrightarrow> Cod f = {}) \<and> inj_on (Fun f) (Dom f)"
thus "section f" using section_if_inj by auto
qed
text\<open>
Section-retraction pairs can also be characterized by an inverse relationship
between the functions they induce.
\<close>
lemma section_retraction_char:
shows "ide (g \<cdot> f) \<longleftrightarrow> antipar f g \<and> compose (Dom f) (Fun g) (Fun f) = (\<lambda>x \<in> Dom f. x)"
- by (metis Fun_comp cod_comp compose_eq' dom_comp ide_char ide_compE seqE)
+ by (metis Fun_comp cod_comp compose_eq' dom_comp ide_char\<^sub>S\<^sub>C ide_compE seqE)
text\<open>
Antiparallel arrows @{term f} and @{term g} are inverses if the functions
they induce are inverses.
\<close>
lemma inverse_arrows_char:
shows "inverse_arrows f g \<longleftrightarrow>
antipar f g \<and> compose (Dom f) (Fun g) (Fun f) = (\<lambda>x \<in> Dom f. x)
\<and> compose (Dom g) (Fun f) (Fun g) = (\<lambda>y \<in> Dom g. y)"
using section_retraction_char by blast
text\<open>
An arrow is an isomorphism if and only if the function it induces is a bijection.
\<close>
lemma iso_char:
shows "iso f \<longleftrightarrow> arr f \<and> bij_betw (Fun f) (Dom f) (Cod f)"
by (metis bij_betw_def image_empty iso_iff_section_and_retraction retraction_char
section_char)
text\<open>
The inverse of an isomorphism is constructed by inverting the induced function.
\<close>
lemma inv_char:
assumes "iso f"
shows "inv f = mkArr (Cod f) (Dom f) (inv_into (Dom f) (Fun f))"
proof -
let ?g = "mkArr (Cod f) (Dom f) (inv_into (Dom f) (Fun f))"
have "ide (f \<cdot> ?g)"
using assms iso_is_retraction retraction_char retraction_if_Img_eq_Cod by simp
moreover have "ide (?g \<cdot> f)"
proof -
let ?g' = "mkArr (Cod f) (Dom f)
(\<lambda>y. if y \<in> Img f then SOME x. x \<in> Dom f \<and> Fun f x = y
else SOME x. x \<in> Dom f)"
have 1: "ide (?g' \<cdot> f)"
using assms iso_is_section section_char section_if_inj by simp
moreover have "?g' = ?g"
proof
show "arr ?g'" using 1 ide_compE by blast
show "\<And>y. y \<in> Cod f \<Longrightarrow> (if y \<in> Img f then SOME x. x \<in> Dom f \<and> Fun f x = y
else SOME x. x \<in> Dom f)
= inv_into (Dom f) (Fun f) y"
proof -
fix y
assume "y \<in> Cod f"
hence "y \<in> Img f" using assms iso_is_retraction retraction_char by metis
thus "(if y \<in> Img f then SOME x. x \<in> Dom f \<and> Fun f x = y
else SOME x. x \<in> Dom f)
= inv_into (Dom f) (Fun f) y"
using inv_into_def by metis
qed
qed
ultimately show ?thesis by auto
qed
ultimately have "inverse_arrows f ?g" by auto
thus ?thesis using inverse_unique by blast
qed
lemma Fun_inv:
assumes "iso f"
shows "Fun (inv f) = restrict (inv_into (Dom f) (Fun f)) (Cod f)"
using assms inv_in_hom inv_char iso_inv_iso iso_is_arr Fun_mkArr by metis
subsection "Monomorphisms and Epimorphisms"
text\<open>
An arrow is a monomorphism if and only if the function it induces is injective.
\<close>
lemma mono_char:
shows "mono f \<longleftrightarrow> arr f \<and> inj_on (Fun f) (Dom f)"
proof
assume f: "mono f"
hence "arr f" using mono_def by auto
moreover have "inj_on (Fun f) (Dom f)"
proof (intro inj_onI)
have 0: "inj_on (S f) (hom unity (dom f))"
proof -
have "hom unity (dom f) \<subseteq> {g. seq f g}"
using f mono_def arrI by auto
hence "\<exists>A. hom unity (dom f) \<subseteq> A \<and> inj_on (S f) A"
using f mono_def by auto
thus ?thesis
by (meson subset_inj_on)
qed
fix x x'
assume x: "x \<in> Dom f" and x': "x' \<in> Dom f" and xx': "Fun f x = Fun f x'"
have "mkPoint (dom f) x \<in> hom unity (dom f) \<and>
mkPoint (dom f) x' \<in> hom unity (dom f)"
using x x' \<open>arr f\<close> mkPoint_in_hom by simp
moreover have "f \<cdot> mkPoint (dom f) x = f \<cdot> mkPoint (dom f) x'"
using \<open>arr f\<close> x x' xx' comp_arr_mkPoint by simp
ultimately have "mkPoint (dom f) x = mkPoint (dom f) x'"
using 0 inj_onD [of "S f" "hom unity (dom f)" "mkPoint (dom f) x"] by simp
thus "x = x'"
using \<open>arr f\<close> x x' img_mkPoint(2) img_mkPoint(2) ide_dom by metis
qed
ultimately show "arr f \<and> inj_on (Fun f) (Dom f)" by auto
next
assume f: "arr f \<and> inj_on (Fun f) (Dom f)"
show "mono f"
proof
show "arr f" using f by auto
show "\<And>g g'. seq f g \<and> seq f g' \<and> f \<cdot> g = f \<cdot> g' \<Longrightarrow> g = g'"
proof -
fix g g'
assume gg': "seq f g \<and> seq f g' \<and> f \<cdot> g = f \<cdot> g'"
show "g = g'"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
show par: "par g g'"
using gg' dom_comp by (metis seqE)
show "Fun g = Fun g'"
proof
fix x
have "x \<notin> Dom g \<Longrightarrow> Fun g x = Fun g' x"
using gg' by (simp add: par Fun_def)
moreover have "x \<in> Dom g \<Longrightarrow> Fun g x = Fun g' x"
proof -
assume x: "x \<in> Dom g"
have "Fun f (Fun g x) = Fun (f \<cdot> g) x"
using gg' x Fun_comp [of f g] by auto
also have "... = Fun f (Fun g' x)"
using par f gg' x monoE by simp
finally have "Fun f (Fun g x) = Fun f (Fun g' x)" by auto
moreover have "Fun g x \<in> Dom f \<and> Fun g' x \<in> Dom f"
using par gg' x Fun_mapsto by fastforce
ultimately show "Fun g x = Fun g' x"
using f gg' inj_onD [of "Fun f" "Dom f" "Fun g x" "Fun g' x"]
by simp
qed
ultimately show "Fun g x = Fun g' x" by auto
qed
qed
qed
qed
qed
text\<open>
Inclusions are monomorphisms.
\<close>
lemma mono_imp_incl:
assumes "incl f"
shows "mono f"
using assms incl_def Fun_incl mono_char by auto
text\<open>
A monomorphism is a section, except in case it has an empty domain set and
a nonempty codomain set.
\<close>
lemma mono_imp_section:
assumes "mono f" and "Dom f = {} \<longrightarrow> Cod f = {}"
shows "section f"
using assms mono_char section_char by auto
text\<open>
An arrow is an epimorphism if and only if either its image coincides with its
codomain, or else the universe has only a single element (in which case all arrows
are epimorphisms).
\<close>
lemma epi_char:
shows "epi f \<longleftrightarrow> arr f \<and> (Img f = Cod f \<or> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t'))"
proof
assume epi: "epi f"
show "arr f \<and> (Img f = Cod f \<or> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t'))"
proof -
have f: "arr f" using epi epi_implies_arr by auto
moreover have "\<not>(\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t') \<Longrightarrow> Img f = Cod f"
proof -
assume "\<not>(\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t')"
from this obtain tt and ff
where B: "tt \<in> Univ \<and> ff \<in> Univ \<and> tt \<noteq> ff" by blast
show "Img f = Cod f"
proof
show "Img f \<subseteq> Cod f" using f Fun_mapsto by auto
show "Cod f \<subseteq> Img f"
proof
let ?g = "mkArr (Cod f) {ff, tt} (\<lambda>y. tt)"
let ?g' = "mkArr (Cod f) {ff, tt} (\<lambda>y. if \<exists>x. x \<in> Dom f \<and> Fun f x = y
then tt else ff)"
let ?b = "mkIde {ff, tt}"
have b: "ide ?b"
by (metis B finite.emptyI finite_imp_setp finite_insert ide_mkIde
insert_subset setp_imp_subset_Univ setp_singleton mem_Collect_eq)
have g: "\<guillemotleft>?g : cod f \<rightarrow> ?b\<guillemotright> \<and> Fun ?g = (\<lambda>y \<in> Cod f. tt)"
using f B in_homI [of ?g "cod f" "mkIde {ff, tt}"] finite_imp_setp by simp
have g': "?g' \<in> hom (cod f) ?b \<and>
Fun ?g' = (\<lambda>y \<in> Cod f. if \<exists>x. x \<in> Dom f \<and> Fun f x = y then tt else ff)"
using f B in_homI [of ?g'] finite_imp_setp by simp
have "?g \<cdot> f = ?g' \<cdot> f"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
show "par (?g \<cdot> f) (?g' \<cdot> f)"
using f g g' by auto
show "Fun (?g \<cdot> f) = Fun (?g' \<cdot> f)"
using f g g' Fun_comp comp_mkArr by fastforce
qed
hence gg': "?g = ?g'"
by (metis (no_types, lifting) epiE epi f g in_homE seqI)
fix y
assume y: "y \<in> Cod f"
have "Fun ?g' y = tt" using gg' g y by simp
hence "(if \<exists>x. x \<in> Dom f \<and> Fun f x = y then tt else ff) = tt"
using g' y by simp
hence "\<exists>x. x \<in> Dom f \<and> Fun f x = y"
using B by argo
thus "y \<in> Img f" by blast
qed
qed
qed
ultimately show "arr f \<and> (Img f = Cod f \<or> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t'))"
by fast
qed
next
show "arr f \<and> (Img f = Cod f \<or> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t')) \<Longrightarrow> epi f"
proof -
have "arr f \<and> Img f = Cod f \<Longrightarrow> epi f"
using retraction_char retraction_is_epi by presburger
moreover have "arr f \<and> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t') \<Longrightarrow> epi f"
proof -
assume f: "arr f \<and> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t')"
have "\<And>f f'. par f f' \<Longrightarrow> f = f'"
proof -
fix f f'
assume ff': "par f f'"
show "f = f'"
- proof (intro arr_eqI)
+ proof (intro arr_eqI\<^sub>S\<^sub>C)
show "par f f'" using ff' by simp
have "\<And>t t'. t \<in> Cod f \<and> t' \<in> Cod f \<Longrightarrow> t = t'"
using f ff' setp_imp_subset_Univ setp_set_ide ide_cod subsetD by blast
thus "Fun f = Fun f'"
using ff' Fun_mapsto [of f] Fun_mapsto [of f']
extensional_arb [of "Fun f" "Dom f"] extensional_arb [of "Fun f'" "Dom f"]
by fastforce
qed
qed
moreover have "\<And>g g'. par (g \<cdot> f) (g' \<cdot> f) \<Longrightarrow> par g g'"
by force
ultimately show "epi f"
using f by (intro epiI; metis)
qed
ultimately show "arr f \<and> (Img f = Cod f \<or> (\<forall>t t'. t \<in> Univ \<and> t' \<in> Univ \<longrightarrow> t = t'))
\<Longrightarrow> epi f"
by auto
qed
qed
text\<open>
An epimorphism is a retraction, except in the case of a degenerate universe with only
a single element.
\<close>
lemma epi_imp_retraction:
assumes "epi f" and "\<exists>t t'. t \<in> Univ \<and> t' \<in> Univ \<and> t \<noteq> t'"
shows "retraction f"
using assms epi_char retraction_char by auto
text\<open>
Retraction/inclusion factorization is unique (not just up to isomorphism -- remember
that the notion of inclusion is not categorical but depends on the arbitrarily chosen
@{term img}).
\<close>
lemma unique_retr_incl_fact:
assumes "seq m e" and "seq m' e'" and "m \<cdot> e = m' \<cdot> e'"
and "incl m" and "incl m'" and "retraction e" and "retraction e'"
shows "m = m'" and "e = e'"
proof -
have 1: "cod m = cod m' \<and> dom e = dom e'"
using assms(1-3) by (metis dom_comp cod_comp)
hence 2: "span e e'" using assms(1-2) by blast
hence 3: "Fun e = Fun e'"
using assms eq_Fun_iff_incl_joinable by meson
hence "img e = img e'" using assms 1 img_def by auto
moreover have "img e = cod e \<and> img e' = cod e'"
using assms(6-7) retraction_char img_def mkIde_set by simp
ultimately have "par e e'" using 2 by simp
- thus "e = e'" using 3 arr_eqI by blast
+ thus "e = e'" using 3 arr_eqI\<^sub>S\<^sub>C by blast
hence "par m m'" using assms(1) assms(2) 1 by fastforce
thus "m = m'" using assms(4) assms(5) incls_coherent by blast
qed
end
section "Concrete Set Categories"
text\<open>
The \<open>set_category\<close> locale is useful for stating results that depend on a
category of @{typ 'a}-sets and functions, without having to commit to a particular
element type @{typ 'a}. However, in applications we often need to work with a
category of sets and functions that is guaranteed to contain sets corresponding
to the subsets of some extrinsically given type @{typ 'a}.
A \emph{concrete set category} is a set category \<open>S\<close> that is equipped
with an injective function @{term \<iota>} from type @{typ 'a} to \<open>S.Univ\<close>.
The following locale serves to facilitate some of the technical aspects of passing
back and forth between elements of type @{typ 'a} and the elements of \<open>S.Univ\<close>.
\<close>
locale concrete_set_category = set_category S setp
for S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and setp :: "'s set \<Rightarrow> bool"
and U :: "'a set"
and \<iota> :: "'a \<Rightarrow> 's" +
assumes UP_mapsto: "\<iota> \<in> U \<rightarrow> Univ"
and inj_UP: "inj_on \<iota> U"
begin
abbreviation UP
where "UP \<equiv> \<iota>"
abbreviation DN
where "DN \<equiv> inv_into U UP"
lemma DN_mapsto:
shows "DN \<in> UP ` U \<rightarrow> U"
by (simp add: inv_into_into)
lemma DN_UP [simp]:
assumes "x \<in> U"
shows "DN (UP x) = x"
using assms inj_UP inv_into_f_f by simp
lemma UP_DN [simp]:
assumes "t \<in> UP ` U"
shows "UP (DN t) = t"
using assms o_def inj_UP by auto
lemma bij_UP:
shows "bij_betw UP U (UP ` U)"
using inj_UP inj_on_imp_bij_betw by blast
lemma bij_DN:
shows "bij_betw DN (UP ` U) U"
using bij_UP bij_betw_inv_into by blast
end
locale replete_concrete_set_category =
replete_set_category S +
concrete_set_category S \<open>\<lambda>A. A \<subseteq> Univ\<close> U UP
for S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and U :: "'a set"
and UP :: "'a \<Rightarrow> 's"
section "Sub-Set Categories"
text\<open>
In this section, we show that a full subcategory of a set category, obtained
by imposing suitable further restrictions on the subsets of the universe
that correspond to objects, is again a set category.
\<close>
locale sub_set_category =
S: set_category +
fixes ssetp :: "'a set \<Rightarrow> bool"
assumes ssetp_singleton: "\<And>t. t \<in> S.Univ \<Longrightarrow> ssetp {t}"
and subset_closed: "\<And>B A. \<lbrakk>B \<subseteq> A; ssetp A\<rbrakk> \<Longrightarrow> ssetp B"
and union_closed: "\<And>A B. \<lbrakk>ssetp A; ssetp B\<rbrakk> \<Longrightarrow> ssetp (A \<union> B)"
and containment: "\<And>A. ssetp A \<Longrightarrow> setp A"
begin
sublocale full_subcategory S \<open>\<lambda>a. S.ide a \<and> ssetp (S.set a)\<close>
by unfold_locales auto
lemma is_full_subcategory:
shows "full_subcategory S (\<lambda>a. S.ide a \<and> ssetp (S.set a))"
..
- lemma ide_char:
+ lemma ide_char\<^sub>S\<^sub>S\<^sub>C:
shows "ide a \<longleftrightarrow> S.ide a \<and> ssetp (S.set a)"
- using ide_char arr_char by fastforce
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by fastforce
- lemma terminal_unity:
+ lemma terminal_unity\<^sub>S\<^sub>S\<^sub>C:
shows "terminal S.unity"
proof
show "ide S.unity"
- using S.terminal_unity S.terminal_def [of S.unity] S.terminal_char2 ide_char
+ using S.terminal_unity\<^sub>S\<^sub>C S.terminal_def [of S.unity] S.terminal_char2 ide_char\<^sub>S\<^sub>S\<^sub>C
ssetp_singleton
by force
thus "\<And>a. ide a \<Longrightarrow> \<exists>!f. in_hom f a S.unity"
- using S.terminal_unity S.terminal_def ide_char ide_char' in_hom_char
+ using S.terminal_unity\<^sub>S\<^sub>C S.terminal_def ide_char\<^sub>S\<^sub>b\<^sub>C ide_char' in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting))
qed
lemma terminal_char:
shows "terminal t \<longleftrightarrow> S.terminal t"
proof
fix t
assume t: "S.terminal t"
have "ide t"
- using t ssetp_singleton ide_char S.terminal_char2 by force
+ using t ssetp_singleton ide_char\<^sub>S\<^sub>S\<^sub>C S.terminal_char2 by force
thus "terminal t"
- using t in_hom_char ide_char arr_char S.terminal_def terminalI by auto
+ using t in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>S\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C S.terminal_def terminalI by auto
next
assume t: "terminal t"
have 1: "S.ide t"
- using t ide_char terminal_def by simp
+ using t ide_char\<^sub>S\<^sub>b\<^sub>C terminal_def by simp
moreover have "card (S.set t) = 1"
proof -
have "card (S.set t) = card (S.hom S.unity t)"
using S.set_def S.inj_img
by (metis 1 S.bij_betw_points_and_set bij_betw_same_card)
also have "... = card (hom S.unity t)"
- using t in_hom_char terminal_def terminal_unity by auto
+ using t in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C terminal_def terminal_unity\<^sub>S\<^sub>S\<^sub>C by auto
also have "... = 1"
proof -
have "\<exists>!f. f \<in> hom S.unity t"
- using t terminal_def terminal_unity by force
+ using t terminal_def terminal_unity\<^sub>S\<^sub>S\<^sub>C by force
moreover have "\<And>A. card A = 1 \<longleftrightarrow> (\<exists>!a. a \<in> A)"
apply (intro iffI)
apply (metis card_1_singletonE empty_iff insert_iff)
using card_1_singleton_iff by auto
ultimately show ?thesis by auto
qed
finally show ?thesis by blast
qed
ultimately show "S.terminal t"
using 1 S.terminal_char1 card_1_singleton_iff
by (metis One_nat_def singleton_iff)
qed
sublocale set_category comp ssetp
proof
text\<open>
Here things are simpler if we define \<open>img\<close> appropriately so that we have
\<open>set = T.set\<close> after accounting for the definition \<open>unity \<equiv> SOME t. terminal t\<close>,
which is different from that of S.unity.
\<close>
have 1: "terminal (SOME t. terminal t)"
- using terminal_unity someI_ex [of terminal] by blast
+ using terminal_unity\<^sub>S\<^sub>S\<^sub>C someI_ex [of terminal] by blast
obtain i where i: "\<guillemotleft>i : S.unity \<rightarrow> SOME t. terminal t\<guillemotright>"
- using terminal_unity someI_ex [of terminal] in_hom_char terminal_def
+ using terminal_unity\<^sub>S\<^sub>S\<^sub>C someI_ex [of terminal] in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C terminal_def
by auto
obtain i' where i': "\<guillemotleft>i' : (SOME t. terminal t) \<rightarrow> S.unity\<guillemotright>"
- using terminal_unity someI_ex [of S.terminal] S.terminal_def
+ using terminal_unity\<^sub>S\<^sub>S\<^sub>C someI_ex [of S.terminal] S.terminal_def
by (metis (no_types, lifting) 1 terminal_def)
have ii': "inverse_arrows i i'"
proof
have "i' \<cdot> i = S.unity"
- using i i' terminal_unity
+ using i i' terminal_unity\<^sub>S\<^sub>S\<^sub>C
by (metis (no_types, lifting) S.comp_in_homI' S.ide_in_hom S.ide_unity S.in_homE
- S.terminalE S.terminal_unity in_hom_char)
+ S.terminalE S.terminal_unity\<^sub>S\<^sub>C in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C)
thus 2: "ide (comp i' i)"
by (metis (no_types, lifting) cod_comp comp_simp i i' ide_char' in_homE seqI')
have "i \<cdot> i' = (SOME t. terminal t)"
using 1
by (metis (no_types, lifting) 2 comp_simp i' ide_compE in_homE inverse_arrowsE
iso_iff_mono_and_retraction point_is_mono retractionI section_retraction_of_iso(2))
thus "ide (comp i i')"
using comp_char
- by (metis (no_types, lifting) 2 ide_char' dom_comp i' ide_compE in_homE seq_char)
+ by (metis (no_types, lifting) 2 ide_char' dom_comp i' ide_compE in_homE seq_char\<^sub>S\<^sub>b\<^sub>C)
qed
interpret set_category_data comp \<open>\<lambda>x. S.some_img (x \<cdot> i)\<close> ..
have i_in_hom: "\<guillemotleft>i : S.unity \<rightarrow> unity\<guillemotright>"
using i unity_def by simp
have i'_in_hom: "\<guillemotleft>i' : unity \<rightarrow> S.unity\<guillemotright>"
using i' unity_def by simp
have "\<And>a. ide a \<Longrightarrow> set a = S.set a"
proof -
fix a
assume a: "ide a"
have "set a = (\<lambda>x. S.some_img (x \<cdot> i)) ` hom unity a"
using set_def by simp
also have "... = S.some_img ` S.hom S.unity a"
proof
show "(\<lambda>x. S.some_img (x \<cdot> i)) ` hom unity a \<subseteq> S.some_img ` S.hom S.unity a"
- using i in_hom_char i_in_hom by auto
+ using i in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C i_in_hom by auto
show "S.some_img ` S.hom S.unity a \<subseteq> (\<lambda>x. S.some_img (x \<cdot> i)) ` hom unity a"
proof
fix b
assume b: "b \<in> S.some_img ` S.hom S.unity a"
obtain x where x: "x \<in> S.hom S.unity a \<and> S.some_img x = b"
using b by blast
have "x \<cdot> i' \<in> hom unity a"
- using x in_hom_char S.comp_in_homI a i' ideD(1) unity_def by force
+ using x in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C S.comp_in_homI a i' ideD(1) unity_def by force
moreover have "S.some_img ((x \<cdot> i') \<cdot> i) = b"
by (metis (no_types, lifting) i ii' x S.comp_assoc calculation comp_simp
ide_compE in_homE inverse_arrowsE mem_Collect_eq S.comp_arr_ide seqI'
- seq_char S.ide_unity unity_def)
+ seq_char\<^sub>S\<^sub>b\<^sub>C S.ide_unity unity_def)
ultimately show "b \<in> (\<lambda>x. S.some_img (x \<cdot> i)) ` hom unity a" by blast
qed
qed
also have "... = S.set a"
using S.set_def by auto
finally show "set a = S.set a" by blast
qed
interpret T: set_category_given_img comp \<open>\<lambda>x. S.some_img (x \<cdot> i)\<close> ssetp
proof
show "Collect terminal \<noteq> {}"
- using terminal_unity by blast
+ using terminal_unity\<^sub>S\<^sub>S\<^sub>C by blast
show "\<And>A' A. \<lbrakk>A' \<subseteq> A; ssetp A\<rbrakk> \<Longrightarrow> ssetp A'"
using subset_closed by blast
show "\<And>A B. \<lbrakk>ssetp A; ssetp B\<rbrakk> \<Longrightarrow> ssetp (A \<union> B)"
using union_closed by simp
show "\<And>A. ssetp A \<Longrightarrow> A \<subseteq> Univ"
using S.setp_imp_subset_Univ containment terminal_char by presburger
show "\<And>a b. \<lbrakk>ide a; ide b; set a = set b\<rbrakk> \<Longrightarrow> a = b"
- using ide_char \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> S.extensional_set by auto
+ using ide_char\<^sub>S\<^sub>b\<^sub>C \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> S.extensional_set by auto
show "\<And>a. ide a \<Longrightarrow> ssetp (set a)"
- using \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> ide_char by force
+ using \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> ide_char\<^sub>S\<^sub>S\<^sub>C by force
show "\<And>A. ssetp A \<Longrightarrow> \<exists>a. ide a \<and> set a = A"
- using S.set_complete \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> containment ide_char by blast
+ using S.set_complete \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> containment ide_char\<^sub>S\<^sub>S\<^sub>C by blast
show "\<And>t. terminal t \<Longrightarrow> t \<in> (\<lambda>x. S.some_img (x \<cdot> i)) ` hom unity t"
using S.set_def S.stable_img \<open>\<And>a. ide a \<Longrightarrow> set a = S.set a\<close> set_def
terminal_char terminal_def
by force
show "\<And>a. ide a \<Longrightarrow> inj_on (\<lambda>x. S.some_img (x \<cdot> i)) (hom unity a)"
proof -
fix a
assume a: "ide a"
show "inj_on (\<lambda>x. S.some_img (x \<cdot> i)) (hom unity a)"
proof
fix x y
assume x: "x \<in> hom unity a" and y: "y \<in> hom unity a"
and eq: "S.some_img (x \<cdot> i) = S.some_img (y \<cdot> i)"
have "x \<cdot> i = y \<cdot> i"
proof -
have "x \<cdot> i \<in> S.hom S.unity a \<and> y \<cdot> i \<in> S.hom S.unity a"
- using in_hom_char \<open>\<guillemotleft>i : S.unity \<rightarrow> unity\<guillemotright>\<close> x y by blast
+ using in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C \<open>\<guillemotleft>i : S.unity \<rightarrow> unity\<guillemotright>\<close> x y by blast
thus ?thesis
- using a eq ide_char S.inj_img [of a] inj_on_def [of S.some_img] by simp
+ using a eq ide_char\<^sub>S\<^sub>b\<^sub>C S.inj_img [of a] inj_on_def [of S.some_img] by simp
qed
have "x = (x \<cdot> i) \<cdot> i'"
by (metis (no_types, lifting) S.comp_arr_ide S.comp_assoc S.inverse_arrowsE
S.match_4 i i' ii' inclusion_preserves_inverse mem_Collect_eq seqI'
- seq_char unity_def x)
+ seq_char\<^sub>S\<^sub>b\<^sub>C unity_def x)
also have "... = (y \<cdot> i) \<cdot> i'"
using \<open>x \<cdot> i = y \<cdot> i\<close> by simp
also have "... = y"
by (metis (no_types, lifting) S.comp_arr_ide S.comp_assoc S.inverse_arrowsE
S.match_4 i i' ii' inclusion_preserves_inverse mem_Collect_eq seqI'
- seq_char unity_def y)
+ seq_char\<^sub>S\<^sub>b\<^sub>C unity_def y)
finally show "x = y" by simp
qed
qed
show "\<And>f f'. \<lbrakk>par f f'; \<And>x. in_hom x unity (dom f) \<Longrightarrow> comp f x = comp f' x\<rbrakk>
\<Longrightarrow> f = f'"
proof -
fix f f'
assume par: "par f f'"
assume eq: "\<And>x. in_hom x unity (dom f) \<Longrightarrow> comp f x = comp f' x"
have "S.par f f'"
- using par arr_char dom_char cod_char by auto
+ using par arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C by auto
moreover have "\<And>x. S.in_hom x S.unity (S.dom f) \<Longrightarrow> f \<cdot> x = f' \<cdot> x"
proof -
fix x
assume x: "S.in_hom x S.unity (S.dom f)"
have "S.in_hom (x \<cdot> i') unity (S.dom f)"
- using i'_in_hom in_hom_char x by blast
+ using i'_in_hom in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C x by blast
hence 1: "in_hom (x \<cdot> i') unity (dom f)"
- using arr_dom dom_simp i in_hom_char par unity_def by force
+ using arr_dom dom_simp i in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C par unity_def by force
hence "comp f (x \<cdot> i') = comp f' (x \<cdot> i')"
using eq by blast
hence "(f \<cdot> (x \<cdot> i')) \<cdot> i = (f' \<cdot> (x \<cdot> i')) \<cdot> i"
using comp_char
by (metis (no_types, lifting) 1 comp_simp in_homE seqI par)
thus "f \<cdot> x = f' \<cdot> x"
by (metis (no_types, lifting) S.comp_arr_dom S.comp_assoc S.comp_inv_arr
- S.in_homE i_in_hom ii' in_hom_char inclusion_preserves_inverse x)
+ S.in_homE i_in_hom ii' in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C inclusion_preserves_inverse x)
qed
ultimately show "f = f'"
using S.extensional_arr by blast
qed
show "\<And>a b F. \<lbrakk>ide a; ide b; F \<in> hom unity a \<rightarrow> hom unity b\<rbrakk>
\<Longrightarrow> \<exists>f. in_hom f a b \<and>
(\<forall>x. in_hom x unity (dom f) \<longrightarrow> comp f x = F x)"
proof -
fix a b F
assume a: "ide a" and b: "ide b" and F: "F \<in> hom unity a \<rightarrow> hom unity b"
have "S.ide a"
- using a ide_char by blast
+ using a ide_char\<^sub>S\<^sub>b\<^sub>C by blast
have "S.ide b"
- using b ide_char by blast
+ using b ide_char\<^sub>S\<^sub>b\<^sub>C by blast
have 1: "(\<lambda>x. F (x \<cdot> i') \<cdot> i) \<in> S.hom S.unity a \<rightarrow> S.hom S.unity b"
proof
fix x
assume x: "x \<in> S.hom S.unity a"
have "x \<cdot> i' \<in> S.hom unity a"
- using i'_in_hom in_hom_char x by blast
+ using i'_in_hom in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C x by blast
hence "x \<cdot> i' \<in> hom unity a"
- using a in_hom_char
- by (metis (no_types, lifting) ideD(1) i'_in_hom in_hom_char mem_Collect_eq)
+ using a in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
+ by (metis (no_types, lifting) ideD(1) i'_in_hom in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C mem_Collect_eq)
hence "F (x \<cdot> i') \<in> hom unity b"
using a b F by blast
hence "F (x \<cdot> i') \<in> S.hom unity b"
- using b in_hom_char by blast
+ using b in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C by blast
thus "F (x \<cdot> i') \<cdot> i \<in> S.hom S.unity b"
- using i in_hom_char unity_def by auto
+ using i in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C unity_def by auto
qed
obtain f where f: "S.in_hom f a b \<and> (\<forall>x. S.in_hom x S.unity (S.dom f)
\<longrightarrow> f \<cdot> x = (\<lambda>x. F (x \<cdot> i') \<cdot> i) x)"
using 1 S.fun_complete_ax \<open>S.ide a\<close> \<open>S.ide b\<close> by presburger
show "\<exists>f. in_hom f a b \<and> (\<forall>x. in_hom x unity (dom f) \<longrightarrow> comp f x = F x)"
proof -
have "in_hom f a b"
- using f in_hom_char ideD(1) a b by presburger
+ using f in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C ideD(1) a b by presburger
moreover have "\<forall>x. in_hom x unity (dom f) \<longrightarrow> comp f x = F x"
proof (intro allI impI)
fix x
assume x: "in_hom x unity (dom f)"
have xi: "S.in_hom (x \<cdot> i) S.unity (S.dom f)"
- using f x i in_hom_char dom_char
+ using f x i in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) in_homE unity_def calculation S.comp_in_homI)
hence 1: "f \<cdot> (x \<cdot> i) = F ((x \<cdot> i) \<cdot> i') \<cdot> i"
using f by blast
hence "((f \<cdot> x) \<cdot> i) \<cdot> i' = (F x \<cdot> i) \<cdot> i'"
by (metis (no_types, lifting) xi S.comp_assoc S.inverse_arrowsE
- S.seqI' i' ii' in_hom_char inclusion_preserves_inverse S.comp_arr_ide)
+ S.seqI' i' ii' in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C inclusion_preserves_inverse S.comp_arr_ide)
hence "f \<cdot> x = F x"
by (metis (no_types, lifting) xi 1 S.invert_side_of_triangle(2) S.match_2 S.match_3
- S.seqI arr_char calculation S.in_homE S.inverse_unique S.isoI
+ S.seqI arr_char\<^sub>S\<^sub>b\<^sub>C calculation S.in_homE S.inverse_unique S.isoI
ii' in_homE inclusion_preserves_inverse)
thus "comp f x = F x"
using comp_char
by (metis (no_types, lifting) comp_simp in_homE seqI calculation x)
qed
ultimately show ?thesis by blast
qed
qed
qed
show "\<exists>img. set_category_given_img comp img ssetp"
using T.set_category_given_img_axioms by blast
qed
lemma is_set_category:
shows "set_category comp ssetp"
..
end
end
diff --git a/thys/Category3/Subcategory.thy b/thys/Category3/Subcategory.thy
--- a/thys/Category3/Subcategory.thy
+++ b/thys/Category3/Subcategory.thy
@@ -1,431 +1,434 @@
(* Title: Subcategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2017
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter "Subcategory"
text\<open>
In this chapter we give a construction of the subcategory of a category
defined by a predicate on arrows subject to closure conditions. The arrows of
the subcategory are directly identified with the arrows of the ambient category.
We also define the related notions of full subcategory and inclusion functor.
\<close>
theory Subcategory
imports Functor
begin
locale subcategory =
C: category C
for C :: "'a comp" (infixr "\<cdot>\<^sub>C" 55)
and Arr :: "'a \<Rightarrow> bool" +
assumes inclusion: "Arr f \<Longrightarrow> C.arr f"
and dom_closed: "Arr f \<Longrightarrow> Arr (C.dom f)"
and cod_closed: "Arr f \<Longrightarrow> Arr (C.cod f)"
and comp_closed: "\<lbrakk> Arr f; Arr g; C.cod f = C.dom g \<rbrakk> \<Longrightarrow> Arr (g \<cdot>\<^sub>C f)"
begin
no_notation C.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation C.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
definition comp (infixr "\<cdot>" 55)
where "g \<cdot> f = (if Arr f \<and> Arr g \<and> C.cod f = C.dom g then g \<cdot>\<^sub>C f else C.null)"
interpretation partial_magma comp
proof
show "\<exists>!n. \<forall>f. n \<cdot> f = n \<and> f \<cdot> n = n"
proof
show 1: "\<forall>f. C.null \<cdot> f = C.null \<and> f \<cdot> C.null = C.null"
by (metis C.comp_null(1) C.ex_un_null comp_def)
show "\<And>n. \<forall>f. n \<cdot> f = n \<and> f \<cdot> n = n \<Longrightarrow> n = C.null"
using 1 C.ex_un_null by metis
qed
qed
lemma null_char [simp]:
shows "null = C.null"
proof -
have "\<forall>f. C.null \<cdot> f = C.null \<and> f \<cdot> C.null = C.null"
by (metis C.comp_null(1) C.ex_un_null comp_def)
thus ?thesis using ex_un_null by (metis comp_null(2))
qed
- lemma ideI:
+ lemma ideI\<^sub>S\<^sub>b\<^sub>C:
assumes "Arr a" and "C.ide a"
shows "ide a"
unfolding ide_def
using assms null_char C.ide_def comp_def by auto
lemma Arr_iff_dom_in_domain:
shows "Arr f \<longleftrightarrow> C.dom f \<in> domains f"
proof
show "C.dom f \<in> domains f \<Longrightarrow> Arr f"
using domains_def comp_def ide_def by fastforce
show "Arr f \<Longrightarrow> C.dom f \<in> domains f"
proof -
assume f: "Arr f"
have "ide (C.dom f)"
using f inclusion C.dom_in_domains C.has_domain_iff_arr C.domains_def
- dom_closed ideI
+ dom_closed ideI\<^sub>S\<^sub>b\<^sub>C
by auto
moreover have "f \<cdot> C.dom f \<noteq> null"
using f comp_def dom_closed null_char inclusion C.comp_arr_dom by force
ultimately show ?thesis
using domains_def by simp
qed
qed
lemma Arr_iff_cod_in_codomain:
shows "Arr f \<longleftrightarrow> C.cod f \<in> codomains f"
proof
show "C.cod f \<in> codomains f \<Longrightarrow> Arr f"
using codomains_def comp_def ide_def by fastforce
show "Arr f \<Longrightarrow> C.cod f \<in> codomains f"
proof -
assume f: "Arr f"
have "ide (C.cod f)"
using f inclusion C.cod_in_codomains C.has_codomain_iff_arr C.codomains_def
- cod_closed ideI
+ cod_closed ideI\<^sub>S\<^sub>b\<^sub>C
by auto
moreover have "C.cod f \<cdot> f \<noteq> null"
using f comp_def cod_closed null_char inclusion C.comp_cod_arr by force
ultimately show ?thesis
using codomains_def by simp
qed
qed
- lemma arr_char:
+ lemma arr_char\<^sub>S\<^sub>b\<^sub>C:
shows "arr f \<longleftrightarrow> Arr f"
proof
show "Arr f \<Longrightarrow> arr f"
using arr_def comp_def Arr_iff_dom_in_domain Arr_iff_cod_in_codomain by auto
show "arr f \<Longrightarrow> Arr f"
proof -
assume f: "arr f"
obtain a where a: "a \<in> domains f \<or> a \<in> codomains f"
using f arr_def by auto
have "f \<cdot> a \<noteq> C.null \<or> a \<cdot> f \<noteq> C.null"
using a domains_def codomains_def null_char by auto
thus "Arr f"
using comp_def by metis
qed
qed
- lemma arrI [intro]:
+ lemma arrI\<^sub>S\<^sub>b\<^sub>C [intro]:
assumes "Arr f"
shows "arr f"
- using assms arr_char by simp
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C by simp
lemma arrE [elim]:
assumes "arr f"
shows "Arr f"
- using assms arr_char by simp
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C by simp
interpretation category comp
using comp_def null_char inclusion comp_closed dom_closed cod_closed
apply unfold_locales
apply auto[1]
- apply (metis Arr_iff_dom_in_domain Arr_iff_cod_in_codomain arr_char arr_def emptyE)
+ apply (metis Arr_iff_dom_in_domain Arr_iff_cod_in_codomain arr_char\<^sub>S\<^sub>b\<^sub>C arr_def emptyE)
proof -
fix f g h
assume gf: "seq g f" and hg: "seq h g"
show 1: "seq (h \<cdot> g) f"
using gf hg inclusion comp_closed comp_def by auto
show "(h \<cdot> g) \<cdot> f = h \<cdot> g \<cdot> f"
- using gf hg 1 C.not_arr_null inclusion comp_def arr_char
+ using gf hg 1 C.not_arr_null inclusion comp_def arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (full_types) C.cod_comp C.comp_assoc)
next
fix f g h
assume hg: "seq h g" and hgf: "seq (h \<cdot> g) f"
show "seq g f"
- using hg hgf comp_def null_char inclusion arr_char comp_closed
+ using hg hgf comp_def null_char inclusion arr_char\<^sub>S\<^sub>b\<^sub>C comp_closed
by (metis (full_types) C.dom_comp)
next
fix f g h
assume hgf: "seq h (g \<cdot> f)" and gf: "seq g f"
show "seq h g"
- using hgf gf comp_def null_char arr_char comp_closed
+ using hgf gf comp_def null_char arr_char\<^sub>S\<^sub>b\<^sub>C comp_closed
by (metis C.seqE C.ext C.match_2)
qed
theorem is_category:
shows "category comp" ..
notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
lemma dom_simp:
assumes "arr f"
shows "dom f = C.dom f"
proof -
have "ide (C.dom f)"
- using assms ideI
- by (meson C.ide_dom arr_char dom_closed inclusion)
+ using assms ideI\<^sub>S\<^sub>b\<^sub>C
+ by (meson C.ide_dom arr_char\<^sub>S\<^sub>b\<^sub>C dom_closed inclusion)
moreover have "f \<cdot> C.dom f \<noteq> null"
- using assms inclusion comp_def null_char dom_closed not_arr_null C.comp_arr_dom arr_char
+ using assms inclusion comp_def null_char dom_closed not_arr_null C.comp_arr_dom arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
ultimately show ?thesis
using dom_eqI ext by blast
qed
- lemma dom_char:
+ lemma dom_char\<^sub>S\<^sub>b\<^sub>C:
shows "dom f = (if arr f then C.dom f else C.null)"
- using dom_simp dom_def arr_def arr_char by auto
+ using dom_simp dom_def arr_def arr_char\<^sub>S\<^sub>b\<^sub>C by auto
lemma cod_simp:
assumes "arr f"
shows "cod f = C.cod f"
proof -
have "ide (C.cod f)"
- using assms ideI
- by (meson C.ide_cod arr_char cod_closed inclusion)
+ using assms ideI\<^sub>S\<^sub>b\<^sub>C
+ by (meson C.ide_cod arr_char\<^sub>S\<^sub>b\<^sub>C cod_closed inclusion)
moreover have "C.cod f \<cdot> f \<noteq> null"
- using assms inclusion comp_def null_char cod_closed not_arr_null C.comp_cod_arr arr_char
+ using assms inclusion comp_def null_char cod_closed not_arr_null C.comp_cod_arr arr_char\<^sub>S\<^sub>b\<^sub>C
by auto
ultimately show ?thesis
using cod_eqI ext by blast
qed
- lemma cod_char:
+ lemma cod_char\<^sub>S\<^sub>b\<^sub>C:
shows "cod f = (if arr f then C.cod f else C.null)"
using cod_simp cod_def arr_def by auto
- lemma in_hom_char:
+ lemma in_hom_char\<^sub>S\<^sub>b\<^sub>C:
shows "\<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<longleftrightarrow> arr a \<and> arr b \<and> arr f \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>"
- using inclusion arr_char cod_closed dom_closed
+ using inclusion arr_char\<^sub>S\<^sub>b\<^sub>C cod_closed dom_closed
by (metis C.arr_iff_in_hom C.in_homE arr_iff_in_hom cod_simp dom_simp in_homE)
- lemma ide_char:
+ lemma ide_char\<^sub>S\<^sub>b\<^sub>C:
shows "ide a \<longleftrightarrow> arr a \<and> C.ide a"
- using ide_in_hom C.ide_in_hom in_hom_char by simp
+ using ide_in_hom C.ide_in_hom in_hom_char\<^sub>S\<^sub>b\<^sub>C by simp
- lemma seq_char:
+ lemma seq_char\<^sub>S\<^sub>b\<^sub>C:
shows "seq g f \<longleftrightarrow> arr f \<and> arr g \<and> C.seq g f"
proof
show "arr f \<and> arr g \<and> C.seq g f \<Longrightarrow> seq g f"
- using arr_char dom_char cod_char by (intro seqI, auto)
+ using arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C by (intro seqI, auto)
show "seq g f \<Longrightarrow> arr f \<and> arr g \<and> C.seq g f"
apply (elim seqE, auto)
- using inclusion arr_char dom_simp cod_simp by auto
+ using inclusion arr_char\<^sub>S\<^sub>b\<^sub>C dom_simp cod_simp by auto
qed
lemma hom_char:
shows "hom a b = C.hom a b \<inter> Collect Arr"
proof
show "hom a b \<subseteq> C.hom a b \<inter> Collect Arr"
- using in_hom_char by auto
+ using in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
show "C.hom a b \<inter> Collect Arr \<subseteq> hom a b"
- using arr_char dom_char cod_char by force
+ using arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C by force
qed
lemma comp_char:
shows "g \<cdot> f = (if arr f \<and> arr g \<and> C.seq g f then g \<cdot>\<^sub>C f else C.null)"
- using arr_char comp_def comp_closed C.ext by force
+ using arr_char\<^sub>S\<^sub>b\<^sub>C comp_def comp_closed C.ext by force
lemma comp_simp:
assumes "seq g f"
shows "g \<cdot> f = g \<cdot>\<^sub>C f"
- using assms comp_char seq_char by metis
+ using assms comp_char seq_char\<^sub>S\<^sub>b\<^sub>C by metis
lemma inclusion_preserves_inverse:
assumes "inverse_arrows f g"
shows "C.inverse_arrows f g"
- using assms ide_char comp_simp arr_char
+ using assms ide_char\<^sub>S\<^sub>b\<^sub>C comp_simp arr_char\<^sub>S\<^sub>b\<^sub>C
by (intro C.inverse_arrowsI, auto)
- lemma iso_char:
+ lemma iso_char\<^sub>S\<^sub>b\<^sub>C:
shows "iso f \<longleftrightarrow> C.iso f \<and> arr f \<and> arr (C.inv f)"
proof
assume f: "iso f"
show "C.iso f \<and> arr f \<and> arr (C.inv f)"
proof -
have 1: "inverse_arrows f (inv f)"
using f inv_is_inverse by auto
have 2: "C.inverse_arrows f (inv f)"
using 1 inclusion_preserves_inverse by blast
moreover have "arr (inv f)"
using 1 iso_is_arr by blast
moreover have "inv f = C.inv f"
using 1 2 C.inv_is_inverse C.inverse_arrow_unique by blast
ultimately show ?thesis using f 2 iso_is_arr by auto
qed
next
assume f: "C.iso f \<and> arr f \<and> arr (C.inv f)"
show "iso f"
proof
have 1: "C.inverse_arrows f (C.inv f)"
using f C.inv_is_inverse by blast
show "inverse_arrows f (C.inv f)"
proof
have 2: "C.inv f \<cdot> f = C.inv f \<cdot>\<^sub>C f \<and> f \<cdot> C.inv f = f \<cdot>\<^sub>C C.inv f"
using f 1 comp_char by fastforce
have 3: "antipar f (C.inv f)"
using f C.seqE seqI dom_simp cod_simp by simp
show "ide (C.inv f \<cdot> f)"
- using 1 2 3 ide_char apply (elim C.inverse_arrowsE) by simp
+ using 1 2 3 ide_char\<^sub>S\<^sub>b\<^sub>C apply (elim C.inverse_arrowsE) by simp
show "ide (f \<cdot> C.inv f)"
- using 1 2 3 ide_char apply (elim C.inverse_arrowsE) by simp
+ using 1 2 3 ide_char\<^sub>S\<^sub>b\<^sub>C apply (elim C.inverse_arrowsE) by simp
qed
qed
qed
- lemma inv_char:
+ lemma inv_char\<^sub>S\<^sub>b\<^sub>C:
assumes "iso f"
shows "inv f = C.inv f"
proof -
have "C.inverse_arrows f (inv f)"
proof
have 1: "inverse_arrows f (inv f)"
- using assms iso_char inv_is_inverse by blast
+ using assms iso_char\<^sub>S\<^sub>b\<^sub>C inv_is_inverse by blast
show "C.ide (inv f \<cdot>\<^sub>C f)"
proof -
have "inv f \<cdot> f = inv f \<cdot>\<^sub>C f"
- using assms 1 inv_in_hom iso_char [of f] comp_char [of "inv f" f] seq_char by auto
+ using assms 1 inv_in_hom iso_char\<^sub>S\<^sub>b\<^sub>C [of f] comp_char [of "inv f" f] seq_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis
- using 1 ide_char arr_char by force
+ using 1 ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by force
qed
show "C.ide (f \<cdot>\<^sub>C inv f)"
proof -
have "f \<cdot> inv f = f \<cdot>\<^sub>C inv f"
- using assms 1 inv_in_hom iso_char [of f] comp_char [of f "inv f"] seq_char by auto
+ using assms 1 inv_in_hom iso_char\<^sub>S\<^sub>b\<^sub>C [of f] comp_char [of f "inv f"] seq_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis
- using 1 ide_char arr_char by force
+ using 1 ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by force
qed
qed
thus ?thesis using C.inverse_arrow_unique C.inv_is_inverse by blast
qed
end
sublocale subcategory \<subseteq> category comp
using is_category by auto
section "Full Subcategory"
locale full_subcategory =
C: category C
for C :: "'a comp"
and Ide :: "'a \<Rightarrow> bool" +
- assumes inclusion: "Ide f \<Longrightarrow> C.ide f"
+ assumes inclusion\<^sub>F\<^sub>S\<^sub>b\<^sub>C: "Ide f \<Longrightarrow> C.ide f"
+ begin
- sublocale full_subcategory \<subseteq> subcategory C "\<lambda>f. C.arr f \<and> Ide (C.dom f) \<and> Ide (C.cod f)"
+ sublocale subcategory C "\<lambda>f. C.arr f \<and> Ide (C.dom f) \<and> Ide (C.cod f)"
by (unfold_locales; simp)
- context full_subcategory
- begin
+ lemma is_subcategory:
+ shows "subcategory C (\<lambda>f. C.arr f \<and> Ide (C.dom f) \<and> Ide (C.cod f))"
+ ..
- lemma in_hom_char:
+ lemma in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C:
shows "\<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<longleftrightarrow> arr a \<and> arr b \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright>"
- using arr_char in_hom_char by auto
+ using arr_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C by auto
text \<open>
Isomorphisms in a full subcategory are inherited from the ambient category.
\<close>
- lemma iso_char:
+ lemma iso_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C:
shows "iso f \<longleftrightarrow> arr f \<and> C.iso f"
proof
assume f: "iso f"
obtain g where g: "inverse_arrows f g" using f by blast
show "arr f \<and> C.iso f"
proof -
have "C.inverse_arrows f g"
using g apply (elim inverse_arrowsE, intro C.inverse_arrowsI)
- using comp_simp ide_char arr_char by auto
+ using comp_simp ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C by auto
thus ?thesis
using f iso_is_arr by blast
qed
next
assume f: "arr f \<and> C.iso f"
obtain g where g: "C.inverse_arrows f g" using f by blast
have "inverse_arrows f g"
proof
show "ide (comp g f)"
using f g
by (metis (no_types, lifting) C.seqE C.ide_compE C.inverse_arrowsE
- arr_char dom_simp ide_dom comp_def)
+ arr_char\<^sub>S\<^sub>b\<^sub>C dom_simp ide_dom comp_def)
show "ide (comp f g)"
using f g C.inverse_arrows_sym [of f g]
by (metis (no_types, lifting) C.seqE C.ide_compE C.inverse_arrowsE
- arr_char dom_simp ide_dom comp_def)
+ arr_char\<^sub>S\<^sub>b\<^sub>C dom_simp ide_dom comp_def)
qed
thus "iso f" by auto
qed
end
section "Inclusion Functor"
text \<open>
If \<open>S\<close> is a subcategory of \<open>C\<close>, then there is an inclusion functor
from \<open>S\<close> to \<open>C\<close>. Inclusion functors are faithful embeddings.
\<close>
locale inclusion_functor =
C: category C +
S: subcategory C Arr
for C :: "'a comp"
and Arr :: "'a \<Rightarrow> bool"
begin
interpretation "functor" S.comp C S.map
- using S.map_def S.arr_char S.inclusion S.dom_char S.cod_char
- S.dom_closed S.cod_closed S.comp_closed S.arr_char S.comp_char
+ using S.map_def S.arr_char\<^sub>S\<^sub>b\<^sub>C S.inclusion S.dom_char\<^sub>S\<^sub>b\<^sub>C S.cod_char\<^sub>S\<^sub>b\<^sub>C
+ S.dom_closed S.cod_closed S.comp_closed S.arr_char\<^sub>S\<^sub>b\<^sub>C S.comp_char
apply unfold_locales
apply auto[4]
by (elim S.seqE, auto)
lemma is_functor:
shows "functor S.comp C S.map" ..
interpretation faithful_functor S.comp C S.map
apply unfold_locales by simp
lemma is_faithful_functor:
shows "faithful_functor S.comp C S.map" ..
interpretation embedding_functor S.comp C S.map
apply unfold_locales by simp
lemma is_embedding_functor:
shows "embedding_functor S.comp C S.map" ..
end
sublocale inclusion_functor \<subseteq> faithful_functor S.comp C S.map
using is_faithful_functor by auto
sublocale inclusion_functor \<subseteq> embedding_functor S.comp C S.map
using is_embedding_functor by auto
text \<open>
The inclusion of a full subcategory is a special case.
Such functors are fully faithful.
\<close>
locale full_inclusion_functor =
C: category C +
S: full_subcategory C Ide
for C :: "'a comp"
and Ide :: "'a \<Rightarrow> bool"
+ begin
- sublocale full_inclusion_functor \<subseteq>
- inclusion_functor C "\<lambda>f. C.arr f \<and> Ide (C.dom f) \<and> Ide (C.cod f)" ..
+ sublocale inclusion_functor C \<open>\<lambda>f. C.arr f \<and> Ide (C.dom f) \<and> Ide (C.cod f)\<close> ..
- context full_inclusion_functor
- begin
+ lemma is_inclusion_functor:
+ shows "inclusion_functor C (\<lambda>f. C.arr f \<and> Ide (C.dom f) \<and> Ide (C.cod f))"
+ ..
interpretation full_functor S.comp C S.map
apply unfold_locales
using S.ide_in_hom
- by (metis (no_types, lifting) C.in_homE S.arr_char S.in_hom_char S.map_simp)
+ by (metis (no_types, lifting) C.in_homE S.arr_char\<^sub>S\<^sub>b\<^sub>C S.in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C S.map_simp)
lemma is_full_functor:
shows "full_functor S.comp C S.map" ..
+ sublocale full_functor S.comp C S.map
+ using is_full_functor by auto
+ sublocale fully_faithful_functor S.comp C S.map ..
+
end
- sublocale full_inclusion_functor \<subseteq> full_functor S.comp C S.map
- using is_full_functor by auto
- sublocale full_inclusion_functor \<subseteq> fully_faithful_functor S.comp C S.map ..
-
end
diff --git a/thys/Category3/ZFC_SetCat.thy b/thys/Category3/ZFC_SetCat.thy
--- a/thys/Category3/ZFC_SetCat.thy
+++ b/thys/Category3/ZFC_SetCat.thy
@@ -1,1012 +1,1009 @@
(* Title: ZFC_SetCat
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2022
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter "ZFC SetCat"
text\<open>
In the statement and proof of the Yoneda Lemma given in theory \<open>Yoneda\<close>,
we sidestepped the issue, of not having a category of ``all'' sets, by axiomatizing
the notion of a ``set category'', showing that for every category we could obtain
a hom-functor into a set category at a higher type, and then proving the Yoneda
lemma for that particular hom-functor. This is perhaps the best we can do within HOL,
because HOL does not provide any type that contains a universe of sets with the closure
properties usually associated with a category \<open>Set\<close> of sets and functions between them.
However, a significant aspect of category theory involves considering ``all''
algebraic structures of a particular kind as the objects of a ``large'' category having
nice closure or completeness properties. Being able to consider a category of sets that
is ``small-complete'', or a cartesian closed category of sets and functions that includes
some infinite sets as objects, are basic examples of this kind of situation.
The purpose of this section is to demonstrate that, although it cannot be done in
pure HOL, if we are willing to accept the existence of a type \<open>V\<close> whose inhabitants
correspond to sets satisfying the axioms of ZFC, then it is possible to construct,
for example, the ``large'' category of sets and functions as it is usually understood
in category theory. Moreover, assuming the existence of such a type is essentially
all we have to do; all the category theory we have developed so far still applies.
Specifically, what we do in this section is to use theory \<open>ZFC_in_HOL\<close>, which provides
an axiomatization of a set-theoretic universe \<open>V\<close>, to construct a ``set category''
\<open>ZFC_SetCat\<close>, whose objects correspond to \<open>V\<close>-sets, whose arrows correspond to functions
between \<open>V\<close>-sets, and which has the small-completeness property traditionally ascribed
to the category of all small sets and functions between them.
\<close>
theory ZFC_SetCat
imports ZFC_in_HOL.ZFC_Cardinals Limit
begin
text\<open>
The following locale constructs the category of classes and functions between them
and shows that it is small complete. The category is obtained simply as the replete
set category at type \<open>V\<close>. This is not yet the category of sets we want, because it
contains objects corresponding to ``large'' \<open>V\<close>-sets.
\<close>
locale ZFC_class_cat
begin
sublocale replete_setcat \<open>undefined :: V\<close> .
lemma admits_small_V_tupling:
assumes "small (I :: V set)"
shows "admits_tupling I"
proof (unfold admits_tupling_def)
let ?\<pi> = "\<lambda>f. UP (VLambda (ZFC_in_HOL.set I) (DN o f))"
have "?\<pi> \<in> (I \<rightarrow> Univ) \<inter> extensional' I \<rightarrow> Univ"
using UP_mapsto by force
moreover have "inj_on ?\<pi> ((I \<rightarrow> Univ) \<inter> extensional' I)"
proof
fix f g
assume f: "f \<in> (I \<rightarrow> Univ) \<inter> extensional' I"
assume g: "g \<in> (I \<rightarrow> Univ) \<inter> extensional' I"
assume eq: "UP (VLambda (ZFC_in_HOL.set I) (DN \<circ> f)) =
UP (VLambda (ZFC_in_HOL.set I) (DN \<circ> g))"
have 1: "VLambda (ZFC_in_HOL.set I) (DN \<circ> f) = VLambda (ZFC_in_HOL.set I) (DN \<circ> g)"
using f g eq
by (meson injD inj_UP)
show "f = g"
proof
fix x
have "x \<notin> I \<Longrightarrow> f x = g x"
using f g extensional'_def [of I] by auto
moreover have "x \<in> I \<Longrightarrow> f x = g x"
proof -
assume x: "x \<in> I"
hence "(DN o f) x = (DN o g) x"
using assms 1 x elts_of_set VLambda_eq_D2 by fastforce
thus "f x = g x"
using f g x comp_apply UP_DN
by (metis IntD1 PiE bij_arr_of bij_betw_imp_surj_on)
qed
ultimately show "f x = g x" by blast
qed
qed
ultimately show "\<exists>\<pi>. \<pi> \<in> (I \<rightarrow> Univ) \<inter> extensional' I \<rightarrow> Univ \<and>
inj_on \<pi> ((I \<rightarrow> Univ) \<inter> extensional' I)"
by blast
qed
corollary admits_small_tupling:
assumes "small I"
shows "admits_tupling I"
proof -
obtain \<phi> where \<phi>: "inj_on \<phi> I \<and> \<phi> ` I \<in> range elts"
using assms small_def by metis
have "admits_tupling (\<phi> ` I)"
using \<phi> admits_small_V_tupling by fastforce
moreover have inv: "bij_betw (inv_into I \<phi>) (\<phi> ` I) I"
by (simp add: \<phi> bij_betw_inv_into inj_on_imp_bij_betw)
ultimately show ?thesis
using admits_tupling_respects_bij by blast
qed
lemma has_small_products:
assumes "small (I :: 'i set)" and "I \<noteq> UNIV"
shows "has_products I"
proof -
have 1: "\<And>I :: V set. small I \<Longrightarrow> has_products I"
using big_UNIV
by (metis admits_small_tupling has_products_iff_admits_tupling)
obtain V_of where V_of: "inj_on V_of I \<and> V_of ` I \<in> range elts"
using assms small_def by metis
have "bij_betw (inv_into I V_of) (V_of ` I) I"
using V_of bij_betw_inv_into bij_betw_imageI by metis
moreover have "small (V_of ` I)"
using assms by auto
ultimately show ?thesis
using assms 1 has_products_preserved_by_bijection by blast
qed
theorem has_small_limits:
assumes "small (UNIV :: 'i set)"
shows "has_limits (undefined :: 'i)"
proof -
have "\<forall>I :: 'i set. I \<noteq> UNIV \<longrightarrow> admits_tupling I"
using assms smaller_than_small subset_UNIV admits_small_tupling by auto
thus ?thesis
using assms has_limits_iff_admits_tupling by blast
qed
end
text\<open>
We now construct the desired category of small sets and functions between them,
as a full subcategory of the category of classes and functions. To show that this
subcategory is small complete, we show that the inclusion creates small products;
that is, a small product of objects corresponding to small sets itself corresponds
to a small set.
\<close>
locale ZFC_set_cat
begin
interpretation Cls: ZFC_class_cat .
+ abbreviation setp
+ where "setp A \<equiv> A \<subseteq> Cls.Univ \<and> small A"
+
sublocale sub_set_category Cls.comp \<open>\<lambda>A. A \<subseteq> Cls.Univ\<close> \<open>\<lambda>A. A \<subseteq> Cls.Univ \<and> small A\<close>
using small_Un smaller_than_small
apply unfold_locales
apply simp_all
apply force
by auto
- abbreviation setp
- where "setp A \<equiv> A \<subseteq> Cls.Univ \<and> small A"
-
- lemma is_set_category:
- shows "set_category comp setp"
- ..
-
- lemma is_full_subcategory:
- shows "full_subcategory Cls.comp (\<lambda>a. Cls.ide a \<and> setp (Cls.set a))"
- using full_subcategory_axioms by blast
+ lemma is_sub_set_category:
+ shows "sub_set_category Cls.comp (\<lambda>A. A \<subseteq> Cls.Univ) (\<lambda>A. A \<subseteq> Cls.Univ \<and> small A)"
+ using sub_set_category_axioms by blast
interpretation incl: full_inclusion_functor Cls.comp \<open>\<lambda>a. Cls.ide a \<and> setp (Cls.set a)\<close>
..
-
+
text\<open>
The following functions establish a bijection between the identities of the category
and the elements of type \<open>V\<close>; which in turn are in bijective correspondence with
small \<open>V\<close>-sets.
\<close>
definition V_of_ide :: "V setcat.arr \<Rightarrow> V"
where "V_of_ide a \<equiv> ZFC_in_HOL.set (Cls.DN ` Cls.set a)"
definition ide_of_V :: "V \<Rightarrow> V setcat.arr"
where "ide_of_V A \<equiv> Cls.mkIde (Cls.UP ` elts A)"
lemma bij_betw_ide_V:
shows "V_of_ide \<in> Collect ide \<rightarrow> UNIV"
and "ide_of_V \<in> UNIV \<rightarrow> Collect ide"
and [simp]: "ide a \<Longrightarrow> ide_of_V (V_of_ide a) = a"
and [simp]: "V_of_ide (ide_of_V A) = A"
and "bij_betw V_of_ide (Collect ide) UNIV"
and "bij_betw ide_of_V UNIV (Collect ide)"
proof -
have "Univ = Cls.Univ"
using terminal_char by presburger
show 1: "V_of_ide \<in> Collect ide \<rightarrow> UNIV"
by blast
show 2: "ide_of_V \<in> UNIV \<rightarrow> Collect ide"
proof
fix A :: V
have "small (elts A)"
by blast
have "Cls.UP ` elts A \<subseteq> Univ \<and> small (Cls.UP ` elts A)"
using Cls.UP_mapsto terminal_char by blast
hence "ide (mkIde (Cls.UP ` elts A))"
using ide_mkIde \<open>Univ = Cls.Univ\<close> by auto
thus "ide_of_V A \<in> Collect ide"
- using ide_of_V_def ide_char
+ using ide_of_V_def ide_char\<^sub>S\<^sub>S\<^sub>C
by (metis (no_types, lifting) Cls.ide_mkIde Cls.set_mkIde arr_mkIde ide_char'
mem_Collect_eq)
qed
show 3: "\<And>a. ide a \<Longrightarrow> ide_of_V (V_of_ide a) = a"
proof -
fix a
assume a: "ide a"
have "ide_of_V (V_of_ide a) =
Cls.mkIde (Cls.UP ` elts (ZFC_in_HOL.set (Cls.DN ` Cls.set a)))"
unfolding ide_of_V_def V_of_ide_def by simp
also have "... = Cls.mkIde (Cls.UP ` Cls.DN ` Cls.set a)"
- using setp_set_ide a ide_char by force
+ using setp_set_ide a ide_char\<^sub>S\<^sub>S\<^sub>C by force
also have "... = Cls.mkIde (Cls.set a)"
proof -
have "Cls.set a \<subseteq> Cls.Univ"
- using a ide_char by blast
+ using a ide_char\<^sub>S\<^sub>S\<^sub>C by blast
hence "Cls.UP ` Cls.DN ` Cls.set a = Cls.set a"
proof -
have "\<And>x. x \<in> Cls.set a \<Longrightarrow> x \<in> Cls.UP ` Cls.DN ` Cls.set a"
using Cls.UP_DN \<open>Cls.set a \<subseteq> Cls.Univ\<close>
by (metis Cls.bij_arr_of bij_betw_def image_inv_into_cancel)
moreover have "\<And>x. x \<in> Cls.UP ` Cls.DN ` Cls.set a \<Longrightarrow> x \<in> Cls.set a"
using \<open>Cls.set a \<subseteq> Cls.Univ\<close>
by (metis Cls.bij_arr_of bij_betw_def image_inv_into_cancel)
ultimately show ?thesis by blast
qed
thus ?thesis by argo
qed
also have "... = a"
- using a Cls.mkIde_set ide_char by blast
+ using a Cls.mkIde_set ide_char\<^sub>S\<^sub>b\<^sub>C by blast
finally show "ide_of_V (V_of_ide a) = a" by simp
qed
show 4: "\<And>A. V_of_ide (ide_of_V A) = A"
proof -
fix A
have "V_of_ide (ide_of_V A) =
ZFC_in_HOL.set (Cls.DN ` Cls.set (Cls.mkIde (Cls.UP ` elts A)))"
unfolding V_of_ide_def ide_of_V_def by simp
also have "... = ZFC_in_HOL.set (Cls.DN ` Cls.UP ` elts A)"
using Cls.set_mkIde [of "Cls.UP ` elts A"] Cls.UP_mapsto by fastforce
also have "... = ZFC_in_HOL.set (elts A)"
using Cls.DN_UP by force
also have "... = A" by simp
finally show "V_of_ide (ide_of_V A) = A" by blast
qed
show "bij_betw V_of_ide (Collect ide) UNIV"
using 1 2 3 4
by (intro bij_betwI) auto
show "bij_betw ide_of_V UNIV (Collect ide)"
using 1 2 3 4
by (intro bij_betwI) blast+
qed
text\<open>
Next, we establish bijections between the hom-sets of the category and certain
subsets of \<open>V\<close> whose elements represent functions.
\<close>
definition V_of_arr :: "V setcat.arr \<Rightarrow> V"
where "V_of_arr f \<equiv> VLambda (V_of_ide (dom f)) (Cls.DN o Cls.Fun f o Cls.UP)"
definition arr_of_V :: "V setcat.arr \<Rightarrow> V setcat.arr \<Rightarrow> V \<Rightarrow> V setcat.arr"
where "arr_of_V a b F \<equiv> Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP o app F o Cls.DN)"
definition vfun
where "vfun A B f \<equiv> f \<in> elts (VPow (vtimes A B)) \<and> elts A = Domain (pairs f) \<and>
single_valued (pairs f)"
lemma small_Collect_vfun:
shows "small (Collect (vfun A B))"
unfolding vfun_def
by (metis (full_types) small_Collect small_elts)
lemma vfunI:
assumes "f \<in> elts A \<rightarrow> elts B"
shows "vfun A B (VLambda A f)"
proof (unfold vfun_def, intro conjI)
show "VLambda A f \<in> elts (VPow (vtimes A B))"
using assms VLambda_def by auto
show "elts A = Domain (pairs (VLambda A f))"
using assms VLambda_def [of A f]
by (metis Domain_fst fst_pairs_VLambda)
show "single_valued (pairs (VLambda A f))"
using assms VLambda_def single_valued_def pairs_iff_elts by fastforce
qed
lemma app_vfun_mapsto:
assumes "vfun A B F"
shows "app F \<in> elts A \<rightarrow> elts B"
proof
have "F \<in> elts (VPow (vtimes A B)) \<and> elts A = Domain (pairs F) \<and> single_valued (pairs F)"
using assms vfun_def by simp
hence 1: "F \<in> elts (VPi A (\<lambda>_. B)) \<and> elts A = Domain (pairs F) \<and> single_valued (pairs F)"
unfolding VPi_def
by (metis (no_types, lifting) down elts_of_set mem_Collect_eq subsetI)
fix x
assume x: "x \<in> elts A"
have "x \<in> Domain (pairs F)"
using assms x vfun_def by blast
thus "app F x \<in> elts B"
using x 1 VPi_D [of F A "\<lambda>_. B" x] by blast
qed
lemma bij_betw_hom_vfun:
shows "V_of_arr \<in> hom a b \<rightarrow> Collect (vfun (V_of_ide a) (V_of_ide b))"
and "\<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> arr_of_V a b \<in> Collect (vfun (V_of_ide a) (V_of_ide b)) \<rightarrow> hom a b"
and "f \<in> hom a b \<Longrightarrow> arr_of_V a b (V_of_arr f) = f"
and "\<lbrakk>ide a; ide b; F \<in> Collect (vfun (V_of_ide a) (V_of_ide b))\<rbrakk>
\<Longrightarrow> V_of_arr (arr_of_V a b F) = F"
and "\<lbrakk>ide a; ide b\<rbrakk>
\<Longrightarrow> bij_betw V_of_arr (hom a b) (Collect (vfun (V_of_ide a) (V_of_ide b)))"
and "\<lbrakk>ide a; ide b\<rbrakk>
\<Longrightarrow> bij_betw (arr_of_V a b) (Collect (vfun (V_of_ide a) (V_of_ide b))) (hom a b)"
proof -
show 1: "\<And>a b. V_of_arr \<in> hom a b \<rightarrow> Collect (vfun (V_of_ide a) (V_of_ide b))"
proof
fix a b f
assume f: "f \<in> hom a b"
have "V_of_arr f = VLambda (V_of_ide (dom f)) (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP)"
unfolding V_of_arr_def by simp
moreover have "vfun (V_of_ide a) (V_of_ide b) ..."
proof -
have "Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP \<in> elts (V_of_ide a) \<rightarrow> elts (V_of_ide b)"
proof
fix x
assume x: "x \<in> elts (V_of_ide a)"
have "(Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) x = Cls.DN (Cls.Fun f (Cls.UP x))"
by simp
moreover have "... \<in> elts (V_of_ide b)"
proof -
have "Cls.UP x \<in> Cls.Dom f"
by (metis (no_types, lifting) Cls.arr_dom_iff_arr Cls.arr_mkIde Cls.set_mkIde
- bij_betw_ide_V(3) arr_char dom_simp f ide_char ide_of_V_def image_eqI
+ bij_betw_ide_V(3) arr_char\<^sub>S\<^sub>b\<^sub>C dom_simp f ide_char\<^sub>S\<^sub>S\<^sub>C ide_of_V_def image_eqI
in_homE mem_Collect_eq x)
hence "Cls.DN (Cls.Fun f (Cls.UP x)) \<in> Cls.DN ` Cls.Cod f"
- using f in_hom_char arr_char Cls.Fun_mapsto [of f] by blast
+ using f in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C Cls.Fun_mapsto [of f] by blast
thus "Cls.DN (Cls.Fun f (Cls.UP x)) \<in> elts (V_of_ide b)"
- unfolding V_of_ide_def
- using f ide_char in_hom_char
- by (metis (no_types, lifting) arr_char cod_simp elts_of_set in_homE
- mem_Collect_eq replacement)
+ by (metis (no_types, lifting) V_of_ide_def arrE cod_char\<^sub>S\<^sub>b\<^sub>C elts_of_set f
+ in_homE mem_Collect_eq replacement)
qed
ultimately show "(Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) x \<in> elts (V_of_ide b)" by argo
qed
thus ?thesis
using f vfunI by blast
qed
ultimately show "V_of_arr f \<in> Collect (vfun (V_of_ide a) (V_of_ide b))" by simp
qed
show 2: "\<And>a b. \<lbrakk>ide a; ide b\<rbrakk>
\<Longrightarrow> arr_of_V a b \<in> Collect (vfun (V_of_ide a) (V_of_ide b)) \<rightarrow> hom a b"
proof -
fix a b
assume a: "ide a" and b: "ide b"
show "arr_of_V a b \<in> Collect (vfun (V_of_ide a) (V_of_ide b)) \<rightarrow> hom a b"
proof
fix F
assume F: "F \<in> Collect (vfun (V_of_ide a) (V_of_ide b))"
have 3: "app F \<in> elts (V_of_ide a) \<rightarrow> elts (V_of_ide b)"
using F app_vfun_mapsto [of "V_of_ide a" "V_of_ide b" F] by blast
have 4: "app F \<in> Cls.DN ` (Cls.set a) \<rightarrow> Cls.DN ` (Cls.set b)"
- using 3 V_of_ide_def a b ide_char by auto
+ using 3 V_of_ide_def a b ide_char\<^sub>S\<^sub>S\<^sub>C by auto
have "arr_of_V a b F = Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN)"
unfolding arr_of_V_def by simp
moreover have "... \<in> hom a b"
proof
show "\<guillemotleft>Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN) : a \<rightarrow> b\<guillemotright>"
proof
have 4: "Cls.arr (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN))"
proof -
have "Cls.set a \<subseteq> Cls.Univ \<and> Cls.set b \<subseteq> Cls.Univ"
- using a b ide_char by blast
+ using a b ide_char\<^sub>S\<^sub>S\<^sub>C by blast
moreover have "Cls.UP \<circ> app F \<circ> Cls.DN \<in> Cls.set a \<rightarrow> Cls.set b"
proof
fix x
assume x: "x \<in> Cls.set a"
have "(Cls.UP \<circ> app F \<circ> Cls.DN) x = Cls.UP (app F (Cls.DN x))"
by simp
moreover have "... \<in> Cls.set b"
by (metis (no_types, lifting) 4 Cls.arr_mkIde Cls.ide_char' Cls.set_mkIde
- PiE V_of_ide_def bij_betw_ide_V(3) b elts_of_set ide_char
+ PiE V_of_ide_def bij_betw_ide_V(3) b elts_of_set ide_char\<^sub>S\<^sub>S\<^sub>C
ide_of_V_def replacement rev_image_eqI x)
ultimately show "(Cls.UP \<circ> app F \<circ> Cls.DN) x \<in> Cls.set b"
by auto
qed
ultimately show ?thesis by blast
qed
show 5: "arr (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN))"
- using a b 4 arr_char ide_char by auto
+ using a b 4 arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C by auto
show "dom (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN)) = a"
- using a 4 5 dom_char ide_char by auto
+ using a 4 5 dom_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C by auto
show "cod (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN)) = b"
- using b 4 5 cod_char ide_char by auto
+ using b 4 5 cod_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>b\<^sub>C by auto
qed
qed
ultimately show "arr_of_V a b F \<in> hom a b" by auto
qed
qed
show 3: "\<And>a b f. f \<in> hom a b \<Longrightarrow> arr_of_V a b (V_of_arr f) = f"
proof -
fix a b f
assume f: "f \<in> hom a b"
have 4: "\<And>x. x \<in> Cls.set a
\<Longrightarrow> (Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x = Cls.Fun f x"
proof -
fix x
assume x: "x \<in> Cls.set a"
have "(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x =
Cls.UP (Cls.DN (Cls.Fun f (Cls.UP (Cls.DN x))))"
by simp
also have "... = Cls.UP (Cls.DN (Cls.Fun f x))"
using x Cls.UP_DN
by (metis (no_types, lifting) Cls.elem_set_implies_incl_in Cls.incl_in_def
Cls.setp_set_ide bij_betw_def replete_setcat.bij_arr_of subset_eq)
also have "... = Cls.Fun f x"
proof -
have "x \<in> Cls.Dom f"
- using x f dom_char by fastforce
+ using x f dom_char\<^sub>S\<^sub>b\<^sub>C by fastforce
hence "Cls.Fun f x \<in> Cls.Cod f"
- using x f Cls.Fun_mapsto in_hom_char by blast
+ using x f Cls.Fun_mapsto in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C by blast
hence "Cls.Fun f x \<in> Cls.Univ"
- using f cod_char in_hom_char
+ using f cod_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) Cls.elem_set_implies_incl_in Cls.incl_in_def
Cls.setp_set_ide subsetD)
thus ?thesis
by (meson bij_betw_inv_into_right replete_setcat.bij_arr_of)
qed
finally show "(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x = Cls.Fun f x"
by blast
qed
have 5: "Cls.arr (Cls.mkArr (Cls.set a) (Cls.set b)
(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN))"
proof -
have "Cls.set a \<subseteq> Cls.Univ \<and> Cls.set b \<subseteq> Cls.Univ"
- using f ide_char codomains_def domains_def in_hom_def by force
+ using f ide_char\<^sub>S\<^sub>b\<^sub>C codomains_def domains_def in_hom_def by force
moreover have "Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN
\<in> Cls.set a \<rightarrow> Cls.set b"
proof
fix x
assume x: "x \<in> Cls.set a"
hence 6: "x \<in> Cls.Dom f"
- using f by (metis (no_types, lifting) dom_char in_homE mem_Collect_eq)
+ using f by (metis (no_types, lifting) dom_char\<^sub>S\<^sub>b\<^sub>C in_homE mem_Collect_eq)
have "(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x = Cls.Fun f x"
using 4 f x by blast
moreover have "... \<in> Cls.Cod f"
using 4 6 f Cls.Fun_mapsto
by (metis (no_types, lifting) Cls.arr_dom_iff_arr Cls.elem_set_implies_incl_in
Cls.ideD(1) Cls.incl_in_def IntE PiE)
moreover have "... = Cls.set b"
- using f by (metis (no_types, lifting) cod_char in_homE mem_Collect_eq)
+ using f by (metis (no_types, lifting) cod_char\<^sub>S\<^sub>b\<^sub>C in_homE mem_Collect_eq)
ultimately show "(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x \<in> Cls.set b"
by auto
qed
ultimately show ?thesis by blast
qed
have "arr_of_V a b (V_of_arr f) =
Cls.mkArr (Cls.set a) (Cls.set b)
(Cls.UP \<circ> app (VLambda (V_of_ide (dom f)) (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP))
\<circ> Cls.DN)"
unfolding arr_of_V_def V_of_arr_def by simp
also have "... = Cls.mkArr (Cls.set a) (Cls.set b)
(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN)"
proof (intro Cls.mkArr_eqI')
show 6: "\<And>x. x \<in> Cls.set a \<Longrightarrow>
(Cls.UP \<circ> app (VLambda (V_of_ide (dom f)) (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP))
\<circ> Cls.DN) x =
(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x"
proof -
fix x
assume x: "x \<in> Cls.set a"
have "(Cls.UP \<circ> app (VLambda (V_of_ide (dom f)) (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP))
\<circ> Cls.DN) x =
Cls.UP (app (VLambda (V_of_ide (dom f)) (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP))
(Cls.DN x))"
by simp
also have "... = (Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x"
proof -
have "Cls.DN x \<in> elts (V_of_ide (dom f))"
using f
- by (metis (no_types, lifting) V_of_ide_def elts_of_set ide_char ide_dom image_eqI
+ by (metis (no_types, lifting) V_of_ide_def elts_of_set ide_char\<^sub>S\<^sub>S\<^sub>C ide_dom image_eqI
in_homE mem_Collect_eq replacement x)
thus ?thesis
using beta by auto
qed
ultimately show "(Cls.UP \<circ> app (VLambda (V_of_ide (dom f))
(Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP))
\<circ> Cls.DN) x =
(Cls.UP \<circ> (Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP) \<circ> Cls.DN) x"
by argo
qed
show "Cls.arr (Cls.mkArr (Cls.set a) (Cls.set b)
(Cls.UP \<circ> app (VLambda (V_of_ide (local.dom f))
(Cls.DN \<circ> Cls.Fun f \<circ> Cls.UP))
\<circ> Cls.DN))"
using 5 6 Cls.mkArr_eqI by auto
qed
also have "... = Cls.mkArr (Cls.set a) (Cls.set b) (Cls.Fun f)"
using 4 5 by force
also have "... = f"
using f Cls.mkArr_Fun
- by (metis (no_types, lifting) arr_char cod_simp dom_char in_homE mem_Collect_eq)
+ by (metis (no_types, lifting) arr_char\<^sub>S\<^sub>b\<^sub>C cod_simp dom_char\<^sub>S\<^sub>b\<^sub>C in_homE mem_Collect_eq)
finally show "arr_of_V a b (V_of_arr f) = f" by blast
qed
show 4: "\<And>a b F. \<lbrakk>ide a; ide b; F \<in> Collect (vfun (V_of_ide a) (V_of_ide b))\<rbrakk>
\<Longrightarrow> V_of_arr (arr_of_V a b F) = F"
proof -
fix a b F
assume a: "ide a" and b: "ide b"
assume F: "F \<in> Collect (vfun (V_of_ide a) (V_of_ide b))"
have "F \<in> elts (VPow (vtimes (V_of_ide a) (V_of_ide b))) \<and>
elts (V_of_ide a) = Domain (pairs F) \<and> single_valued (pairs F)"
using F vfun_def by simp
hence 5: "F \<in> elts (VPi (V_of_ide a) (\<lambda>_. V_of_ide b))"
unfolding VPi_def
by (metis (no_types, lifting) down elts_of_set mem_Collect_eq subsetI)
let ?f = "Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \<circ> app F \<circ> Cls.DN)"
have 6: "Cls.arr ?f"
proof -
have "Cls.set a \<subseteq> Cls.Univ \<and> Cls.set b \<subseteq> Cls.Univ"
- using a b ide_char by blast
+ using a b ide_char\<^sub>S\<^sub>S\<^sub>C by blast
moreover have "Cls.UP \<circ> app F \<circ> Cls.DN \<in> Cls.set a \<rightarrow> Cls.set b"
proof
fix x
assume x: "x \<in> Cls.set a"
have "(Cls.UP \<circ> app F \<circ> Cls.DN) x = Cls.UP (app F (Cls.DN x))"
by simp
moreover have "... \<in> Cls.set b"
proof -
have "app F (Cls.DN x) \<in> Cls.DN ` Cls.set b"
- using a b ide_char x F app_vfun_mapsto [of "V_of_ide a" "V_of_ide b" F] V_of_ide_def
+ using a b ide_char\<^sub>S\<^sub>S\<^sub>C x F app_vfun_mapsto [of "V_of_ide a" "V_of_ide b" F]
+ V_of_ide_def
by auto
thus ?thesis
using \<open>Cls.set a \<subseteq> Cls.Univ \<and> Cls.set b \<subseteq> Cls.Univ\<close>
by (metis Cls.bij_arr_of bij_betw_def imageI image_inv_into_cancel)
qed
ultimately show "(Cls.UP \<circ> app F \<circ> Cls.DN) x \<in> Cls.set b" by auto
qed
ultimately show ?thesis by blast
qed
have "V_of_arr (arr_of_V a b F) = VLambda (V_of_ide (dom ?f)) (Cls.DN \<circ> Cls.Fun ?f \<circ> Cls.UP)"
unfolding V_of_arr_def arr_of_V_def by simp
also have "... = VLambda (V_of_ide a) (Cls.DN \<circ> Cls.Fun ?f \<circ> Cls.UP)"
unfolding V_of_ide_def
- using a b 6 ide_char V_of_ide_def dom_char Cls.dom_mkArr arrI by auto
+ using a b 6 ide_char\<^sub>S\<^sub>S\<^sub>C V_of_ide_def dom_char\<^sub>S\<^sub>b\<^sub>C Cls.dom_mkArr arrI\<^sub>S\<^sub>b\<^sub>C by auto
also have "... = VLambda (V_of_ide a)
(Cls.DN \<circ>
restrict (Cls.UP \<circ> app F \<circ> Cls.DN) (Cls.set a) \<circ> Cls.UP)"
using 6 Cls.Fun_mkArr by simp
also have "... = VLambda (V_of_ide a) (app F)"
proof -
have 7: "\<And>x. x \<in> elts (V_of_ide a) \<Longrightarrow>
(Cls.DN \<circ> restrict (Cls.UP \<circ> app F \<circ> Cls.DN) (Cls.set a) \<circ> Cls.UP) x =
app F x"
unfolding V_of_ide_def
using a
apply simp
- by (metis (no_types, lifting) Cls.bij_arr_of a bij_betw_def empty_iff ide_char
+ by (metis (no_types, lifting) Cls.bij_arr_of a bij_betw_def empty_iff ide_char\<^sub>S\<^sub>S\<^sub>C
image_eqI image_inv_into_cancel)
have 8: "\<And>x. x \<in> elts (V_of_ide a) \<Longrightarrow>
(Cls.DN \<circ> restrict (Cls.UP \<circ> app F \<circ> Cls.DN) (Cls.set a) \<circ> Cls.UP) x
\<in> elts (V_of_ide b)"
using 5 7 VPi_D by fastforce
have "VLambda (V_of_ide a) (app F) \<in> elts (VPi (V_of_ide a) (\<lambda>_. V_of_ide b))"
using 5 VPi_I VPi_D by auto
moreover have "VLambda (V_of_ide a)
(Cls.DN \<circ>
restrict (Cls.UP \<circ> app F \<circ> Cls.DN) (Cls.set a) \<circ> Cls.UP)
\<in> elts (VPi (V_of_ide a) (\<lambda>_. V_of_ide b))"
using 8 VPi_I by auto
moreover have "\<And>x. x \<in> elts (V_of_ide a) \<Longrightarrow>
app (VLambda (V_of_ide a)
(Cls.DN \<circ>
restrict (Cls.UP \<circ> app F \<circ> Cls.DN) (Cls.set a) \<circ>
Cls.UP)) x =
app F x"
using 7 beta by auto
ultimately show ?thesis
using fun_ext by simp
qed
also have "... = F"
using 5 eta [of F "V_of_ide a" "\<lambda>_. V_of_ide b"] by auto
finally show "V_of_arr (arr_of_V a b F) = F" by blast
qed
show "\<lbrakk>ide a; ide b\<rbrakk>
\<Longrightarrow> bij_betw V_of_arr (hom a b) (Collect (vfun (V_of_ide a) (V_of_ide b)))"
using 1 2 3 4
apply (intro bij_betwI)
apply blast
apply blast
by auto
show "\<lbrakk>ide a; ide b\<rbrakk>
\<Longrightarrow> bij_betw (arr_of_V a b) (Collect (vfun (V_of_ide a) (V_of_ide b))) (hom a b)"
using 1 2 3 4
apply (intro bij_betwI)
apply blast
apply blast
by auto
qed
lemma small_hom:
shows "small (hom a b)"
proof (cases "ide a \<and> ide b")
assume 1: "\<not> (ide a \<and> ide b)"
have "\<And>f. \<not> \<guillemotleft>f : a \<rightarrow> b\<guillemotright>"
using 1 in_hom_def ide_cod ide_dom by blast
hence "hom a b = {}"
by blast
thus ?thesis by simp
next
assume 1: "ide a \<and> ide b"
show ?thesis
using 1 bij_betw_hom_vfun(5) small_Collect_vfun small_def
by (metis (no_types, lifting) bij_betw_def small_iff_range)
qed
text\<open>
We can now show that the inclusion of the subcategory into the ambient category \<open>Cls\<close>
creates small products. To do this, we consider a product in \<open>Cls\<close> of objects of the
subcategory indexed by a small set \<open>I\<close>. Since \<open>Cls\<close> is a replete set category,
by a previous result we know that the elements of a product object \<open>p\<close> in \<open>Cls\<close>
correspond to its points; that is, to the elements of \<open>hom unity p\<close>.
The elements of \<open>hom unity p\<close> in turn correspond to \<open>I\<close>-tuples. By carrying out
the construction of the set of \<open>I\<close>-tuples in \<open>V\<close> and exploiting the bijections between
homs of the subcatgory and \<open>V\<close>-sets, we can obtain an injection of \<open>hom unity p\<close>
to the extension of a \<open>V\<close>-set, thus showing \<open>hom unity p\<close> is small.
Since \<open>hom unity p\<close> is small, it determines an object of the subcategory,
which must then be a product in the subcategory, in view of the fact that the
subcategory is full.
\<close>
lemma has_small_V_products:
assumes "small (I :: V set)"
shows "has_products I"
proof (unfold has_products_def, intro conjI impI allI)
show "I \<noteq> UNIV"
using assms big_UNIV by blast
fix J D
assume D: "discrete_diagram J comp D \<and> Collect (partial_magma.arr J) = I"
interpret J: category J
using D discrete_diagram_def by blast
interpret D: discrete_diagram J comp D
using D by blast
interpret incloD: composite_functor J comp Cls.comp D map ..
interpret incloD: discrete_diagram J Cls.comp incloD.map
using D.is_discrete
by unfold_locales auto
interpret incloD: diagram_in_set_category J Cls.comp \<open>\<lambda>A. A \<subseteq> Cls.Univ\<close> incloD.map
..
have 1: "small (Collect J.ide)"
using assms D D.is_discrete by argo
show "\<exists>a. has_as_product J D a"
proof -
have 2: "\<exists>a. Cls.has_as_product J incloD.map a"
proof -
have "Collect J.ide \<noteq> UNIV"
using J.ide_def by blast
thus ?thesis
using 1 D.is_discrete Cls.has_small_products [of "Collect J.ide"]
Cls.has_products_def [of "Collect J.ide"] incloD.discrete_diagram_axioms
by presburger
qed
obtain \<Pi>D where \<Pi>D: "Cls.has_as_product J incloD.map \<Pi>D"
using 2 by blast
interpret \<Pi>D: constant_functor J Cls.comp \<Pi>D
using \<Pi>D Cls.product_is_ide
by unfold_locales auto
obtain \<pi> where \<pi>: "product_cone J Cls.comp incloD.map \<Pi>D \<pi>"
using \<Pi>D Cls.has_as_product_def by blast
interpret \<pi>: product_cone J Cls.comp incloD.map \<Pi>D \<pi>
using \<pi> by blast
have "small (Cls.hom Cls.unity \<Pi>D)"
proof -
obtain \<phi> where \<phi>: "bij_betw \<phi> (Cls.hom Cls.unity \<Pi>D) (incloD.cones Cls.unity)"
using incloD.limits_are_sets_of_cones \<pi>.limit_cone_axioms by blast
let ?J = "ZFC_in_HOL.set (Collect J.arr)"
let ?V_of_point = "\<lambda>x. VLambda ?J (\<lambda>j. V_of_arr (\<phi> x j))"
let ?Tuples = "VPi ?J (\<lambda>j. ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))"
have V_of_point: "?V_of_point \<in> Cls.hom Cls.unity \<Pi>D \<rightarrow> elts ?Tuples"
proof
fix x
assume x: "x \<in> Cls.hom Cls.unity \<Pi>D"
have \<phi>x: "\<phi> x \<in> incloD.cones Cls.unity"
using \<phi> x
unfolding bij_betw_def by blast
interpret \<phi>x: cone J Cls.comp incloD.map Cls.unity \<open>\<phi> x\<close>
using \<phi>x by blast
have "\<And>j. J.arr j \<Longrightarrow> V_of_arr (\<phi> x j)
\<in> elts (ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))"
proof -
fix j
assume j: "J.arr j"
have "\<phi> x j \<in> hom Cls.unity (D j)"
by (metis (mono_tags, lifting) D.preserves_ide \<phi>x.component_in_hom cod_simp
- ideD(1,3) in_hom_char incloD.is_discrete incloD.preserves_cod j map_simp
- mem_Collect_eq o_apply terminal_char2 terminal_unity)
+ ideD(1,3) in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C incloD.is_discrete incloD.preserves_cod j map_simp
+ mem_Collect_eq o_apply terminal_char2 terminal_unity\<^sub>S\<^sub>S\<^sub>C)
moreover have "V_of_arr ` hom Cls.unity (D j) =
elts (ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))"
using small_hom replacement by simp
ultimately show "V_of_arr (\<phi> x j)
\<in> elts (ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))"
using j \<phi>x bij_betw_hom_vfun(1) by blast
qed
thus "?V_of_point x \<in> elts ?Tuples"
using VPi_I by fastforce
qed
have "?V_of_point ` Cls.hom Cls.unity \<Pi>D \<in> range elts"
proof -
have "?V_of_point ` Cls.hom Cls.unity \<Pi>D \<subseteq> elts ?Tuples"
using V_of_point by blast
thus ?thesis
using smaller_than_small down_raw by auto
qed
moreover have "inj_on ?V_of_point (Cls.hom Cls.unity \<Pi>D)"
proof
fix x y
assume x: "x \<in> Cls.hom Cls.unity \<Pi>D" and y: "y \<in> Cls.hom Cls.unity \<Pi>D"
and eq: "?V_of_point x = ?V_of_point y"
have \<phi>x: "\<phi> x \<in> incloD.cones Cls.unity"
using \<phi> x
unfolding bij_betw_def by blast
have \<phi>y: "\<phi> y \<in> incloD.cones Cls.unity"
using \<phi> y
unfolding bij_betw_def by blast
interpret \<phi>x: cone J Cls.comp incloD.map Cls.unity \<open>\<phi> x\<close>
using \<phi>x by blast
interpret \<phi>y: cone J Cls.comp incloD.map Cls.unity \<open>\<phi> y\<close>
using \<phi>y by blast
have "\<phi> x = \<phi> y"
proof -
have "\<And>j. j \<in> elts ?J \<Longrightarrow> \<phi> x j = \<phi> y j"
proof -
fix j
assume j: "j \<in> elts ?J"
hence 3: "J.arr j"
by (simp add: 1 incloD.is_discrete)
have 4: "ide (D j)"
using 3 incloD.is_discrete D.preserves_ide by force
have 5:"ide Cls.unity"
- using Cls.terminal_unity terminal_char terminal_def by auto
+ using Cls.terminal_unity\<^sub>S\<^sub>C terminal_char terminal_def by auto
have \<phi>xj: "\<phi> x j \<in> hom Cls.unity (D j)"
- using 3 4 5 incloD.is_discrete \<phi>x.preserves_hom \<phi>x.A.map_simp in_hom_char
+ using 3 4 5 incloD.is_discrete \<phi>x.preserves_hom \<phi>x.A.map_simp in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) J.ide_char \<phi>x.component_in_hom ideD(1) map_simp
mem_Collect_eq o_apply)
have \<phi>yj: "\<phi> y j \<in> hom Cls.unity (D j)"
- using 3 4 5 incloD.is_discrete \<phi>y.preserves_hom \<phi>x.A.map_simp in_hom_char
+ using 3 4 5 incloD.is_discrete \<phi>y.preserves_hom \<phi>x.A.map_simp in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) J.ide_char \<phi>y.component_in_hom ideD(1) map_simp
mem_Collect_eq o_apply)
show "\<phi> x j = \<phi> y j"
proof -
have "\<And>j. j \<in> elts ?J \<Longrightarrow> V_of_arr (\<phi> x j) = V_of_arr (\<phi> y j)"
using x eq VLambda_eq_D2 by blast
thus ?thesis
using V_of_arr_def
by (metis (mono_tags, lifting) j \<phi>xj \<phi>yj bij_betw_hom_vfun(3) mem_Collect_eq)
qed
qed
moreover have "elts ?J = Collect J.arr"
by (simp add: 1 incloD.is_discrete)
ultimately show ?thesis
using \<phi>x.is_extensional \<phi>y.is_extensional
by (metis HOL.ext mem_Collect_eq)
qed
thus "x = y"
using x y \<phi> bij_betw_imp_inj_on inj_on_def
by (metis (no_types, lifting))
qed
ultimately show "small (Cls.hom Cls.unity \<Pi>D)"
using small_def by blast
qed
hence "small (Cls.set \<Pi>D)"
by (simp add: Cls.set_def)
hence 2: "ide \<Pi>D"
- using ide_char Cls.setp_set_ide Cls.product_is_ide \<Pi>D by blast
+ using ide_char\<^sub>S\<^sub>S\<^sub>C Cls.setp_set_ide Cls.product_is_ide \<Pi>D by blast
interpret \<Pi>D': constant_functor J comp \<Pi>D
using 2 by unfold_locales
interpret \<pi>': cone J comp D \<Pi>D \<pi>
proof -
have "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>\<pi> j : \<Pi>D \<rightarrow> D j\<guillemotright>"
proof
fix j
assume j: "J.arr j"
show 3: "arr (\<pi> j)"
by (metis (mono_tags, lifting) 2 D.as_nat_trans.preserves_cod D.is_discrete
- D.preserves_ide \<pi>.component_in_hom ideD(1) ideD(3) in_homE in_hom_char
+ D.preserves_ide \<pi>.component_in_hom ideD(1) ideD(3) in_homE in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
j map_simp o_apply)
show "dom (\<pi> j) = \<Pi>D"
- using 3 arr_char dom_char \<pi>.preserves_dom by auto
+ using 3 arr_char\<^sub>S\<^sub>b\<^sub>C dom_char\<^sub>S\<^sub>b\<^sub>C \<pi>.preserves_dom by auto
show "cod (\<pi> j) = D j"
- using 3 arr_char cod_char \<pi>.preserves_cod
+ using 3 arr_char\<^sub>S\<^sub>b\<^sub>C cod_char\<^sub>S\<^sub>b\<^sub>C \<pi>.preserves_cod
by (metis (no_types, lifting) Cls.ideD(3) D.preserves_arr functor.preserves_ide
incloD.is_discrete incloD.is_functor incloD.preserves_cod j map_simp o_apply)
qed
moreover have "D.mkCone \<pi> = \<pi>"
using \<pi>.is_extensional null_char by auto
ultimately show "cone J comp D \<Pi>D \<pi>"
using 2 D.cone_mkCone [of \<Pi>D \<pi>] by simp
qed
interpret \<pi>': product_cone J comp D \<Pi>D \<pi>
proof
fix a \<chi>'
assume \<chi>': "D.cone a \<chi>'"
interpret \<chi>': cone J comp D a \<chi>'
using \<chi>' by blast
have a: "Cls.ide a"
- using ide_char \<chi>'.A.value_is_ide by blast
+ using ide_char\<^sub>S\<^sub>b\<^sub>C \<chi>'.A.value_is_ide by blast
moreover have "\<And>j. J.arr j \<Longrightarrow> Cls.in_hom (\<chi>' j) a (incloD.map j)"
proof -
fix j
assume j: "J.arr j"
have "\<guillemotleft>\<chi>' j : a \<rightarrow> D j\<guillemotright>"
using j D.is_discrete \<chi>'.component_in_hom by force
thus "Cls.in_hom (\<chi>' j) a (incloD.map j)"
- using a j D.is_discrete in_hom_char map_simp by auto
+ using a j D.is_discrete in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C map_simp by auto
qed
ultimately have 3: "incloD.cone a (incloD.mkCone \<chi>')"
using incloD.cone_mkCone [of a \<chi>'] by blast
interpret \<chi>: cone J Cls.comp incloD.map a \<open>incloD.mkCone \<chi>'\<close>
using 3 by blast
have univ: "\<exists>!f. Cls.in_hom f a \<Pi>D \<and> incloD.cones_map f \<pi> = incloD.mkCone \<chi>'"
using \<chi>.cone_axioms \<pi>.is_universal [of a "incloD.mkCone \<chi>'"] by blast
have 4: "incloD.mkCone \<chi>' = \<chi>'"
using D.as_nat_trans.preserves_reflects_arr D.preserves_arr Limit.cone_def
\<chi>' \<chi>'.is_extensional identity_functor.intro identity_functor.map_def
incloD.as_nat_trans.is_extensional o_apply
by fastforce
have 5: "D.mkCone \<pi> = \<pi>"
using \<pi>.is_extensional null_char by auto
have 6: "\<And>f. Cls.in_hom f a \<Pi>D \<Longrightarrow> incloD.cones_map f \<pi> = D.cones_map f \<pi>"
proof -
fix f
assume f: "Cls.in_hom f a \<Pi>D"
have "incloD.cones_map f \<pi> = (\<lambda>j. if J.arr j then Cls.comp (\<pi> j) f else Cls.null)"
using f \<pi>.cone_axioms by auto
also have "... = (\<lambda>j. if J.arr j then comp (\<pi> j) f else null)"
proof -
have "\<And>j. J.arr j \<Longrightarrow> Cls.comp (\<pi> j) f = comp (\<pi> j) f"
- using f 2 comp_char in_hom_char seq_char
+ using f 2 comp_char in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C seq_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) Cls.ext Cls.in_homE \<chi>'.ide_apex
- \<pi>'.preserves_reflects_arr arr_char ide_char)
+ \<pi>'.preserves_reflects_arr arr_char\<^sub>S\<^sub>b\<^sub>C ide_char\<^sub>S\<^sub>S\<^sub>C)
thus ?thesis
using null_char by auto
qed
also have "... = D.cones_map f \<pi>"
proof -
have "\<pi> \<in> D.cones (cod f)"
proof -
have "\<guillemotleft>f : a \<rightarrow> \<Pi>D\<guillemotright>"
- using f 2 in_hom_char \<chi>'.ide_apex ideD(1) by presburger
+ using f 2 in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C \<chi>'.ide_apex ideD(1) by presburger
thus ?thesis
using f \<pi>'.cone_axioms by blast
qed
thus ?thesis
using \<open>\<pi> \<in> D.cones (cod f)\<close> by simp
qed
finally show "incloD.cones_map f \<pi> = D.cones_map f \<pi>" by blast
qed
moreover have "\<And>f. \<guillemotleft>f : a \<rightarrow> \<Pi>D\<guillemotright> \<Longrightarrow> Cls.in_hom f a \<Pi>D"
- using in_hom_char by blast
+ using in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C by blast
show "\<exists>!f. \<guillemotleft>f : a \<rightarrow> \<Pi>D\<guillemotright> \<and> D.cones_map f \<pi> = \<chi>'"
proof -
have "\<exists>f. \<guillemotleft>f : a \<rightarrow> \<Pi>D\<guillemotright> \<and> D.cones_map f \<pi> = \<chi>'"
- using 2 4 5 6 univ \<chi>'.ide_apex ideD(1) in_hom_char
+ using 2 4 5 6 univ \<chi>'.ide_apex ideD(1) in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting))
moreover have "\<And>f g. \<lbrakk>\<guillemotleft>f : a \<rightarrow> \<Pi>D\<guillemotright> \<and> D.cones_map f \<pi> = \<chi>';
\<guillemotleft>g : a \<rightarrow> \<Pi>D\<guillemotright> \<and> D.cones_map g \<pi> = \<chi>'\<rbrakk>
\<Longrightarrow> f = g"
- using 2 4 5 6 univ \<chi>'.ide_apex ideD(1) in_hom_char by auto
+ using 2 4 5 6 univ \<chi>'.ide_apex ideD(1) in_hom_char\<^sub>F\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show ?thesis by blast
qed
qed
show ?thesis
using \<pi>'.product_cone_axioms has_as_product_def by blast
qed
qed
corollary has_small_products:
assumes "small I" and "I \<noteq> UNIV"
shows "has_products I"
proof -
have 1: "\<And>I :: V set. small I \<Longrightarrow> has_products I"
using has_small_V_products by blast
obtain \<phi> where \<phi>: "inj_on \<phi> I \<and> \<phi> ` I \<in> range elts"
using assms small_def by metis
have "bij_betw (inv_into I \<phi>) (\<phi> ` I) I"
using \<phi> bij_betw_inv_into bij_betw_imageI by metis
moreover have "small (\<phi> ` I)"
using assms by auto
ultimately show ?thesis
using assms 1 has_products_preserved_by_bijection by blast
qed
theorem has_small_limits:
assumes "category (J :: 'j comp)" and "small (Collect (partial_magma.arr J))"
shows "has_limits_of_shape J"
proof -
interpret J: category J
using assms by blast
have "small (Collect J.ide)"
using assms smaller_than_small [of "Collect J.arr" "Collect J.ide"] by fastforce
moreover have "Collect J.ide \<noteq> UNIV"
using J.ide_def by blast
moreover have "Collect J.arr \<noteq> UNIV"
using J.not_arr_null by blast
ultimately show "has_limits_of_shape J"
using assms has_small_products has_limits_if_has_products [of J] by blast
qed
sublocale concrete_set_category comp setp UNIV Cls.UP
proof
show "Cls.UP \<in> UNIV \<rightarrow> Univ"
using Cls.UP_mapsto terminal_char by presburger
show "inj Cls.UP"
using Cls.inj_UP by blast
qed
lemma is_concrete_set_category:
shows "concrete_set_category comp setp UNIV Cls.UP"
..
end
text\<open>
In pure HOL (without ZFC), we were able to show that every category \<open>C\<close> has a ``hom functor'',
but there was necessarily a dependence of the target set category of the hom functor
on the arrow type of \<open>C\<close>. Using the construction of the present theory, we can now show
that every ``locally small'' category \<open>C\<close> has a hom functor, whose target is the same set
category for all such \<open>C\<close>. To obtain such a hom functor requires a choice, for each hom-set
\<open>hom a b\<close> of \<open>C\<close>, of an injection of \<open>hom a b\<close> to the extension of a \<open>V\<close>-set.
\<close>
locale locally_small_category =
category +
assumes locally_small: "\<lbrakk>ide a; ide b\<rbrakk> \<Longrightarrow> small (hom b a)"
begin
interpretation Cop: dual_category C ..
interpretation CopxC: product_category Cop.comp C ..
interpretation S: ZFC_set_cat .
definition Hom
where "Hom \<equiv> \<lambda>(b, a). S.UP o (SOME \<phi>. \<phi> ` hom b a \<in> range elts \<and> inj_on \<phi> (hom b a))"
interpretation Hom: hom_functor C S.comp S.setp Hom
proof
have 1: "\<And>a b. Hom (b, a) \<in> hom b a \<rightarrow> S.Univ \<and> inj_on (Hom (b, a)) (hom b a)"
proof -
fix a b
show "Hom (b, a) \<in> hom b a \<rightarrow> S.Univ \<and> inj_on (Hom (b, a)) (hom b a)"
proof (cases "ide a \<and> ide b")
show "\<not> (ide a \<and> ide b) \<Longrightarrow> ?thesis"
using inj_on_def by fastforce
assume ab: "ide a \<and> ide b"
show ?thesis
proof
have 1: "\<exists>\<phi>. \<phi> ` hom b a \<in> range elts \<and> inj_on \<phi> (hom b a)"
using ab locally_small [of a b] small_def [of "hom b a"] by blast
let ?\<phi> = "SOME \<phi>. \<phi> ` hom b a \<in> range elts \<and> inj_on \<phi> (hom b a)"
have \<phi>: "?\<phi> ` hom b a \<in> range elts \<and> inj_on ?\<phi> (hom b a)"
using 1 someI_ex [of "\<lambda>\<phi>. \<phi> ` hom b a \<in> range elts \<and> inj_on \<phi> (hom b a)"]
by blast
show "Hom (b, a) \<in> hom b a \<rightarrow> S.Univ"
unfolding Hom_def
using \<phi> S.UP_mapsto by auto
show "inj_on (Hom (b, a)) (hom b a)"
unfolding Hom_def
apply simp
using ab \<phi> S.inj_UP comp_inj_on injD inj_on_def
by (metis (no_types, lifting))
qed
qed
qed
show "\<And>f. arr f \<Longrightarrow> Hom (dom f, cod f) f \<in> S.Univ"
using 1 by blast
show "\<And>b a. \<lbrakk>ide b; ide a\<rbrakk> \<Longrightarrow> inj_on (Hom (b, a)) (hom b a)"
using 1 by blast
show "\<And>b a. \<lbrakk>ide b; ide a\<rbrakk> \<Longrightarrow> Hom (b, a) ` hom b a \<subseteq> S.S.Univ \<and> small (Hom (b, a) ` hom b a)"
using 1 locally_small S.terminal_char by force
qed
lemma has_ZFC_hom_functor:
shows "hom_functor C S.comp S.setp Hom"
..
text\<open>
Using this result, we can now state a more traditional version of the Yoneda Lemma
in which the target category of the Yoneda functor is the same for all locally small
categories.
\<close>
interpretation Y: yoneda_functor C S.comp S.setp Hom
..
theorem ZFC_yoneda_lemma:
assumes "ide a" and "functor Cop.comp S.comp F"
shows "\<exists>\<phi>. bij_betw \<phi> (S.set (F a)) {\<tau>. natural_transformation Cop.comp S.comp (Y.Y a) F \<tau>}"
proof -
interpret F: "functor" Cop.comp S.comp F
using assms(2) by blast
interpret F: set_valued_functor Cop.comp S.comp S.setp F
..
interpret Ya: yoneda_functor_fixed_object C S.comp S.setp Hom a
using assms(1) by unfold_locales blast
interpret Ya: yoneda_lemma C S.comp S.setp Hom F a
..
show ?thesis
using Ya.yoneda_lemma by blast
qed
end
+ interpretation ZFCSetCat: ZFC_set_cat .
+
end
diff --git a/thys/MonoidalCategory/FreeMonoidalCategory.thy b/thys/MonoidalCategory/FreeMonoidalCategory.thy
--- a/thys/MonoidalCategory/FreeMonoidalCategory.thy
+++ b/thys/MonoidalCategory/FreeMonoidalCategory.thy
@@ -1,3473 +1,3473 @@
(* Title: FreeMonoidalCategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2017
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter "The Free Monoidal Category"
text_raw\<open>
\label{fmc-chap}
\<close>
theory FreeMonoidalCategory
imports Category3.Subcategory MonoidalFunctor
begin
text \<open>
In this theory, we use the monoidal language of a category @{term C} defined in
@{theory MonoidalCategory.MonoidalCategory} to give a construction of the free monoidal category
\<open>\<F>C\<close> generated by @{term C}.
The arrows of \<open>\<F>C\<close> are the equivalence classes of formal arrows obtained
by declaring two formal arrows to be equivalent if they are parallel and have the
same diagonalization.
Composition, tensor, and the components of the associator and unitors are all
defined in terms of the corresponding syntactic constructs.
After defining \<open>\<F>C\<close> and showing that it does indeed have the structure of
a monoidal category, we prove the freeness: every functor from @{term C} to a
monoidal category @{term D} extends uniquely to a strict monoidal functor from
\<open>\<F>C\<close> to @{term D}.
We then consider the full subcategory \<open>\<F>\<^sub>SC\<close> of \<open>\<F>C\<close> whose objects
are the equivalence classes of diagonal identity terms
({\em i.e.}~equivalence classes of lists of identity arrows of @{term C}),
and we show that this category is monoidally equivalent to \<open>\<F>C\<close>.
In addition, we show that \<open>\<F>\<^sub>SC\<close> is the free strict monoidal category,
as any functor from \<open>C\<close> to a strict monoidal category @{term D} extends uniquely
to a strict monoidal functor from \<open>\<F>\<^sub>SC\<close> to @{term D}.
\<close>
section "Syntactic Construction"
locale free_monoidal_category =
monoidal_language C
for C :: "'c comp"
begin
no_notation C.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation C.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
text \<open>
Two terms of the monoidal language of @{term C} are defined to be equivalent if
they are parallel formal arrows with the same diagonalization.
\<close>
abbreviation equiv
where "equiv t u \<equiv> Par t u \<and> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
text \<open>
Arrows of \<open>\<F>C\<close> will be the equivalence classes of formal arrows
determined by the relation @{term equiv}. We define here the property of being an
equivalence class of the relation @{term equiv}. Later we show that this property
coincides with that of being an arrow of the category that we will construct.
\<close>
type_synonym 'a arr = "'a term set"
definition ARR where "ARR f \<equiv> f \<noteq> {} \<and> (\<forall>t. t \<in> f \<longrightarrow> f = Collect (equiv t))"
lemma not_ARR_empty:
shows "\<not>ARR {}"
using ARR_def by simp
lemma ARR_eqI:
assumes "ARR f" and "ARR g" and "f \<inter> g \<noteq> {}"
shows "f = g"
using assms ARR_def by fastforce
text \<open>
We will need to choose a representative of each equivalence class as a normal form.
The requirements we have of these representatives are: (1) the normal form of an
arrow @{term t} is equivalent to @{term t}; (2) equivalent arrows have identical
normal forms; (3) a normal form is a canonical term if and only if its diagonalization
is an identity. It follows from these properties and coherence that a term and its
normal form have the same evaluation in any monoidal category. We choose here as a
normal form for an arrow @{term t} the particular term @{term "Inv (Cod t\<^bold>\<down>) \<^bold>\<cdot> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<cdot> Dom t\<^bold>\<down>"}.
However, the only specific properties of this definition we actually use are the
three we have just stated.
\<close>
definition norm ("\<^bold>\<parallel>_\<^bold>\<parallel>")
where "\<^bold>\<parallel>t\<^bold>\<parallel> = Inv (Cod t\<^bold>\<down>) \<^bold>\<cdot> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<cdot> Dom t\<^bold>\<down>"
text \<open>
If @{term t} is a formal arrow, then @{term t} is equivalent to its normal form.
\<close>
lemma equiv_norm_Arr:
assumes "Arr t"
shows "equiv \<^bold>\<parallel>t\<^bold>\<parallel> t"
proof -
have "Par t (Inv (Cod t\<^bold>\<down>) \<^bold>\<cdot> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<cdot> Dom t\<^bold>\<down>)"
using assms Diagonalize_in_Hom red_in_Hom Inv_in_Hom Arr_implies_Ide_Dom
Arr_implies_Ide_Cod Ide_implies_Arr Can_red
by auto
moreover have "\<^bold>\<lfloor>(Inv (Cod t\<^bold>\<down>) \<^bold>\<cdot> \<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<cdot> Dom t\<^bold>\<down>)\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide
Diagonalize_in_Hom Diagonalize_Inv [of "Cod t\<^bold>\<down>"] Diag_Diagonalize
CompDiag_Diag_Dom [of "\<^bold>\<lfloor>t\<^bold>\<rfloor>"] CompDiag_Cod_Diag [of "\<^bold>\<lfloor>t\<^bold>\<rfloor>"]
by (simp add: Diagonalize_red [of "Cod t"] Can_red(1))
ultimately show ?thesis using norm_def by simp
qed
text \<open>
Equivalent arrows have identical normal forms.
\<close>
lemma norm_respects_equiv:
assumes "equiv t u"
shows "\<^bold>\<parallel>t\<^bold>\<parallel> = \<^bold>\<parallel>u\<^bold>\<parallel>"
using assms norm_def by simp
text \<open>
The normal form of an arrow is canonical if and only if its diagonalization
is an identity term.
\<close>
lemma Can_norm_iff_Ide_Diagonalize:
assumes "Arr t"
shows "Can \<^bold>\<parallel>t\<^bold>\<parallel> \<longleftrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms norm_def Can_implies_Arr Arr_implies_Ide_Dom Arr_implies_Ide_Cod Can_red
Inv_preserves_Can Diagonalize_preserves_Can red_in_Hom Diagonalize_in_Hom
Ide_Diagonalize_Can
by fastforce
text \<open>
We now establish various additional properties of normal forms that are consequences
of the three already proved. The definition \<open>norm_def\<close> is not used subsequently.
\<close>
lemma norm_preserves_Can:
assumes "Can t"
shows "Can \<^bold>\<parallel>t\<^bold>\<parallel>"
using assms Can_implies_Arr Can_norm_iff_Ide_Diagonalize Ide_Diagonalize_Can by simp
lemma Par_Arr_norm:
assumes "Arr t"
shows "Par \<^bold>\<parallel>t\<^bold>\<parallel> t"
using assms equiv_norm_Arr by auto
lemma Diagonalize_norm [simp]:
assumes "Arr t"
shows " \<^bold>\<lfloor>\<^bold>\<parallel>t\<^bold>\<parallel>\<^bold>\<rfloor> = \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms equiv_norm_Arr by auto
lemma unique_norm:
assumes "ARR f"
shows "\<exists>!t. \<forall>u. u \<in> f \<longrightarrow> \<^bold>\<parallel>u\<^bold>\<parallel> = t"
proof
have 1: "(SOME t. t \<in> f) \<in> f"
using assms ARR_def someI_ex [of "\<lambda>t. t \<in> f"] by auto
show "\<And>t. \<forall>u. u \<in> f \<longrightarrow> \<^bold>\<parallel>u\<^bold>\<parallel> = t \<Longrightarrow> t = \<^bold>\<parallel>SOME t. t \<in> f\<^bold>\<parallel>"
using assms ARR_def 1 by auto
show "\<forall>u. u \<in> f \<longrightarrow> \<^bold>\<parallel>u\<^bold>\<parallel> = \<^bold>\<parallel>SOME t. t \<in> f\<^bold>\<parallel>"
using assms ARR_def 1 norm_respects_equiv by blast
qed
lemma Dom_norm:
assumes "Arr t"
shows "Dom \<^bold>\<parallel>t\<^bold>\<parallel> = Dom t"
using assms Par_Arr_norm by metis
lemma Cod_norm:
assumes "Arr t"
shows "Cod \<^bold>\<parallel>t\<^bold>\<parallel> = Cod t"
using assms Par_Arr_norm by metis
lemma norm_in_Hom:
assumes "Arr t"
shows "\<^bold>\<parallel>t\<^bold>\<parallel> \<in> Hom (Dom t) (Cod t)"
using assms Par_Arr_norm [of t] by simp
text \<open>
As all the elements of an equivalence class have the same normal form, we can
use the normal form of an arbitrarily chosen element as a canonical representative.
\<close>
definition rep where "rep f \<equiv> \<^bold>\<parallel>SOME t. t \<in> f\<^bold>\<parallel>"
lemma rep_in_ARR:
assumes "ARR f"
shows "rep f \<in> f"
using assms ARR_def someI_ex [of "\<lambda>t. t \<in> f"] equiv_norm_Arr rep_def ARR_def
by fastforce
lemma Arr_rep_ARR:
assumes "ARR f"
shows "Arr (rep f)"
using assms ARR_def rep_in_ARR by auto
text \<open>
We next define a function \<open>mkarr\<close> that maps formal arrows to their equivalence classes.
For terms that are not formal arrows, the function yields the empty set.
\<close>
definition mkarr where "mkarr t = Collect (equiv t)"
lemma mkarr_extensionality:
assumes "\<not>Arr t"
shows "mkarr t = {}"
using assms mkarr_def by simp
lemma ARR_mkarr:
assumes "Arr t"
shows "ARR (mkarr t)"
using assms ARR_def mkarr_def by auto
lemma mkarr_memb_ARR:
assumes "ARR f" and "t \<in> f"
shows "mkarr t = f"
using assms ARR_def mkarr_def by simp
lemma mkarr_rep_ARR [simp]:
assumes "ARR f"
shows "mkarr (rep f) = f"
using assms rep_in_ARR mkarr_memb_ARR by auto
lemma Arr_in_mkarr:
assumes "Arr t"
shows "t \<in> mkarr t"
using assms mkarr_def by simp
text \<open>
Two terms are related by @{term equiv} iff they are both formal arrows
and have identical normal forms.
\<close>
lemma equiv_iff_eq_norm:
shows "equiv t u \<longleftrightarrow> Arr t \<and> Arr u \<and> \<^bold>\<parallel>t\<^bold>\<parallel> = \<^bold>\<parallel>u\<^bold>\<parallel>"
proof
show "equiv t u \<Longrightarrow> Arr t \<and> Arr u \<and> \<^bold>\<parallel>t\<^bold>\<parallel> = \<^bold>\<parallel>u\<^bold>\<parallel>"
using mkarr_def Arr_in_mkarr ARR_mkarr unique_norm by blast
show "Arr t \<and> Arr u \<and> \<^bold>\<parallel>t\<^bold>\<parallel> = \<^bold>\<parallel>u\<^bold>\<parallel> \<Longrightarrow> equiv t u"
using Par_Arr_norm Diagonalize_norm by metis
qed
lemma norm_norm [simp]:
assumes "Arr t"
shows "\<^bold>\<parallel>\<^bold>\<parallel>t\<^bold>\<parallel>\<^bold>\<parallel> = \<^bold>\<parallel>t\<^bold>\<parallel>"
proof -
have "t \<in> mkarr t"
using assms Arr_in_mkarr by blast
moreover have "\<^bold>\<parallel>t\<^bold>\<parallel> \<in> mkarr t"
using assms equiv_norm_Arr mkarr_def by simp
ultimately show ?thesis using assms ARR_mkarr unique_norm by auto
qed
lemma norm_in_ARR:
assumes "ARR f" and "t \<in> f"
shows "\<^bold>\<parallel>t\<^bold>\<parallel> \<in> f"
using assms ARR_def equiv_iff_eq_norm norm_norm Par_Arr_norm by fastforce
lemma norm_rep_ARR [simp]:
assumes "ARR f"
shows "\<^bold>\<parallel>rep f\<^bold>\<parallel> = rep f"
using assms ARR_def someI_ex [of "\<lambda>t. t \<in> f"] rep_def norm_norm by fastforce
lemma norm_memb_eq_rep_ARR:
assumes "ARR f" and "t \<in> f"
shows "norm t = rep f"
using assms ARR_def someI_ex [of "\<lambda>t. t \<in> f"] unique_norm rep_def by metis
lemma rep_mkarr:
assumes "Arr f"
shows "rep (mkarr f) = \<^bold>\<parallel>f\<^bold>\<parallel>"
using assms ARR_mkarr Arr_in_mkarr norm_memb_eq_rep_ARR by fastforce
text \<open>
To prove that two terms determine the same equivalence class,
it suffices to show that they are parallel formal arrows with
identical diagonalizations.
\<close>
lemma mkarr_eqI [intro]:
assumes "Par f g" and "\<^bold>\<lfloor>f\<^bold>\<rfloor> = \<^bold>\<lfloor>g\<^bold>\<rfloor>"
shows "mkarr f = mkarr g"
using assms by (metis ARR_mkarr equiv_iff_eq_norm rep_mkarr mkarr_rep_ARR)
text \<open>
We use canonical representatives to lift the formal domain and codomain functions
from terms to equivalence classes.
\<close>
abbreviation DOM where "DOM f \<equiv> Dom (rep f)"
abbreviation COD where "COD f \<equiv> Cod (rep f)"
lemma DOM_mkarr:
assumes "Arr t"
shows "DOM (mkarr t) = Dom t"
using assms rep_mkarr by (metis Par_Arr_norm)
lemma COD_mkarr:
assumes "Arr t"
shows "COD (mkarr t) = Cod t"
using assms rep_mkarr by (metis Par_Arr_norm)
text \<open>
A composition operation can now be defined on equivalence classes
using the syntactic constructor \<open>Comp\<close>.
\<close>
definition comp (infixr "\<cdot>" 55)
where "comp f g \<equiv> (if ARR f \<and> ARR g \<and> DOM f = COD g
then mkarr ((rep f) \<^bold>\<cdot> (rep g)) else {})"
text \<open>
We commence the task of showing that the composition \<open>comp\<close> so defined
determines a category.
\<close>
interpretation partial_magma comp
apply unfold_locales
using comp_def not_ARR_empty by metis
notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text \<open>
The empty set serves as the null for the composition.
\<close>
lemma null_char:
shows "null = {}"
proof -
let ?P = "\<lambda>n. \<forall>f. f \<cdot> n = n \<and> n \<cdot> f = n"
have "?P {}" using comp_def not_ARR_empty by simp
moreover have "\<exists>!n. ?P n" using ex_un_null by metis
ultimately show ?thesis using null_def theI_unique [of ?P "{}"]
by (metis comp_null(2))
qed
lemma ARR_comp:
assumes "ARR f" and "ARR g" and "DOM f = COD g"
shows "ARR (f \<cdot> g)"
using assms comp_def Arr_rep_ARR ARR_mkarr(1) by simp
lemma DOM_comp [simp]:
assumes "ARR f" and "ARR g" and "DOM f = COD g"
shows "DOM (f \<cdot> g) = DOM g"
using assms comp_def ARR_comp Arr_rep_ARR DOM_mkarr by simp
lemma COD_comp [simp]:
assumes "ARR f" and "ARR g" and "DOM f = COD g"
shows "COD (f \<cdot> g) = COD f"
using assms comp_def ARR_comp Arr_rep_ARR COD_mkarr by simp
lemma comp_assoc:
assumes "g \<cdot> f \<noteq> null" and "h \<cdot> g \<noteq> null"
shows "h \<cdot> (g \<cdot> f) = (h \<cdot> g) \<cdot> f"
proof -
have 1: "ARR f \<and> ARR g \<and> ARR h \<and> DOM h = COD g \<and> DOM g = COD f"
using assms comp_def not_ARR_empty null_char by metis
hence 2: "Arr (rep f) \<and> Arr (rep g) \<and> Arr (rep h) \<and>
Dom (rep h) = Cod (rep g) \<and> Dom (rep g) = Cod (rep f)"
using Arr_rep_ARR by simp
have 3: "h \<cdot> g \<cdot> f = mkarr (rep h \<^bold>\<cdot> rep (mkarr (rep g \<^bold>\<cdot> rep f)))"
using 1 comp_def ARR_comp COD_comp by simp
also have "... = mkarr (rep h \<^bold>\<cdot> rep g \<^bold>\<cdot> rep f)"
proof -
have "equiv (rep h \<^bold>\<cdot> rep (mkarr (rep g \<^bold>\<cdot> rep f))) (rep h \<^bold>\<cdot> rep g \<^bold>\<cdot> rep f)"
proof -
have "Par (rep h \<^bold>\<cdot> rep (mkarr (rep g \<^bold>\<cdot> rep f))) (rep h \<^bold>\<cdot> rep g \<^bold>\<cdot> rep f)"
using 1 2 3 DOM_mkarr ARR_comp COD_comp mkarr_extensionality not_ARR_empty
by (metis Arr.simps(4) Cod.simps(4) Dom.simps(4) snd_map_prod)
(* Here metis claims it is not using snd_map_prod, but removing it fails. *)
moreover have "\<^bold>\<lfloor>rep h \<^bold>\<cdot> rep (mkarr (rep g \<^bold>\<cdot> rep f))\<^bold>\<rfloor> = \<^bold>\<lfloor>rep h \<^bold>\<cdot> rep g \<^bold>\<cdot> rep f\<^bold>\<rfloor>"
using 1 2 Arr_rep_ARR rep_mkarr rep_in_ARR assms(1) ARR_comp mkarr_extensionality
comp_def equiv_iff_eq_norm norm_memb_eq_rep_ARR null_char
by auto
ultimately show ?thesis using equiv_iff_eq_norm by blast
qed
thus ?thesis
using mkarr_def by force
qed
also have "... = mkarr ((rep h \<^bold>\<cdot> rep g) \<^bold>\<cdot> rep f)"
proof -
have "Par (rep h \<^bold>\<cdot> rep g \<^bold>\<cdot> rep f) ((rep h \<^bold>\<cdot> rep g) \<^bold>\<cdot> rep f)"
using 2 by simp
moreover have "\<^bold>\<lfloor>rep h \<^bold>\<cdot> rep g \<^bold>\<cdot> rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>(rep h \<^bold>\<cdot> rep g) \<^bold>\<cdot> rep f\<^bold>\<rfloor>"
using 2 Diag_Diagonalize by (simp add: CompDiag_assoc)
ultimately show ?thesis
using equiv_iff_eq_norm by (simp add: mkarr_def)
qed
also have "... = mkarr (rep (mkarr (rep h \<^bold>\<cdot> rep g)) \<^bold>\<cdot> rep f)"
proof -
have "equiv (rep (mkarr (rep h \<^bold>\<cdot> rep g)) \<^bold>\<cdot> rep f) ((rep h \<^bold>\<cdot> rep g) \<^bold>\<cdot> rep f)"
proof -
have "Par (rep (mkarr (rep h \<^bold>\<cdot> rep g)) \<^bold>\<cdot> rep f) ((rep h \<^bold>\<cdot> rep g) \<^bold>\<cdot> rep f)"
using 1 2 Arr_rep_ARR DOM_comp ARR_comp COD_comp comp_def by auto
moreover have "\<^bold>\<lfloor>rep (mkarr (rep h \<^bold>\<cdot> rep g)) \<^bold>\<cdot> rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>(rep h \<^bold>\<cdot> rep g) \<^bold>\<cdot> rep f\<^bold>\<rfloor>"
using assms(2) 1 2 ARR_comp Arr_rep_ARR mkarr_extensionality rep_mkarr rep_in_ARR
equiv_iff_eq_norm norm_memb_eq_rep_ARR comp_def null_char
by simp
ultimately show ?thesis using equiv_iff_eq_norm by blast
qed
thus ?thesis
using mkarr_def by auto
qed
also have "... = (h \<cdot> g) \<cdot> f"
using 1 comp_def ARR_comp DOM_comp by simp
finally show ?thesis by blast
qed
lemma Comp_in_comp_ARR:
assumes "ARR f" and "ARR g" and "DOM f = COD g"
and "t \<in> f" and "u \<in> g"
shows "t \<^bold>\<cdot> u \<in> f \<cdot> g"
proof -
have "equiv (t \<^bold>\<cdot> u) (rep f \<^bold>\<cdot> rep g)"
proof -
have 1: "Par (t \<^bold>\<cdot> u) (rep f \<^bold>\<cdot> rep g)"
using assms ARR_def Arr_rep_ARR COD_mkarr DOM_mkarr mkarr_memb_ARR
mkarr_extensionality
by (metis (no_types, lifting) Arr.simps(4) Cod.simps(4) Dom.simps(4) snd_map_prod)
(* Here metis claims it is not using snd_map_prod, but removing it fails. *)
moreover have "\<^bold>\<lfloor>t \<^bold>\<cdot> u\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f \<^bold>\<cdot> rep g\<^bold>\<rfloor>"
using assms 1 rep_in_ARR equiv_iff_eq_norm norm_memb_eq_rep_ARR
by (metis (no_types, lifting) Arr.simps(4) Diagonalize.simps(4))
ultimately show ?thesis by simp
qed
thus ?thesis
using assms comp_def mkarr_def by simp
qed
text \<open>
Ultimately, we will show that that the identities of the category are those
equivalence classes, all of whose members diagonalize to formal identity arrows,
having the further property that their canonical representative is a formal
endo-arrow.
\<close>
definition IDE where "IDE f \<equiv> ARR f \<and> (\<forall>t. t \<in> f \<longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>) \<and> DOM f = COD f"
lemma IDE_implies_ARR:
assumes "IDE f"
shows "ARR f"
using assms IDE_def ARR_def by auto
lemma IDE_mkarr_Ide:
assumes "Ide a"
shows "IDE (mkarr a)"
proof -
have "DOM (mkarr a) = COD (mkarr a)"
using assms mkarr_def equiv_iff_eq_norm Par_Arr_norm COD_mkarr DOM_mkarr Ide_in_Hom
by (metis Ide_implies_Can Inv_Ide Ide_implies_Arr Inv_preserves_Can(2))
moreover have "ARR (mkarr a) \<and> (\<forall>t. t \<in> mkarr a \<longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>)"
proof -
have "ARR (mkarr a)" using assms ARR_mkarr Ide_implies_Arr by simp
moreover have "\<forall>t. t \<in> mkarr a \<longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>"
using assms mkarr_def Diagonalize_preserves_Ide by fastforce
ultimately show ?thesis by blast
qed
ultimately show ?thesis using IDE_def by blast
qed
lemma IDE_implies_ide:
assumes "IDE a"
shows "ide a"
proof (unfold ide_def)
have "a \<cdot> a \<noteq> null"
proof -
have "rep a \<^bold>\<cdot> rep a \<in> a \<cdot> a"
using assms IDE_def comp_def Arr_rep_ARR Arr_in_mkarr by simp
thus ?thesis
using null_char by auto
qed
moreover have "\<And>f. (f \<cdot> a \<noteq> null \<longrightarrow> f \<cdot> a = f) \<and> (a \<cdot> f \<noteq> null \<longrightarrow> a \<cdot> f = f)"
proof
fix f :: "'c arr"
show "a \<cdot> f \<noteq> null \<longrightarrow> a \<cdot> f = f"
proof
assume f: "a \<cdot> f \<noteq> null"
hence "ARR f"
using comp_def null_char by auto
have "rep a \<^bold>\<cdot> rep f \<in> a \<cdot> f"
using assms f Comp_in_comp_ARR comp_def rep_in_ARR null_char by metis
moreover have "rep a \<^bold>\<cdot> rep f \<in> f"
proof -
have "rep f \<in> f"
using \<open>ARR f\<close> rep_in_ARR by auto
moreover have "equiv (rep a \<^bold>\<cdot> rep f) (rep f)"
proof -
have 1: "Par (rep a \<^bold>\<cdot> rep f) (rep f)"
using assms f comp_def mkarr_extensionality Arr_rep_ARR IDE_def null_char
by (metis Cod.simps(4) Dom.simps(4))
moreover have "\<^bold>\<lfloor>rep a \<^bold>\<cdot> rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms f 1 comp_def IDE_def CompDiag_Ide_Diag Diag_Diagonalize(1)
Diag_Diagonalize(2) Diag_Diagonalize(3) rep_in_ARR
by auto
ultimately show ?thesis by auto
qed
ultimately show ?thesis
using \<open>ARR f\<close> ARR_def by auto
qed
ultimately show "a \<cdot> f = f"
using mkarr_memb_ARR comp_def by auto
qed
show "f \<cdot> a \<noteq> null \<longrightarrow> f \<cdot> a = f"
proof
assume f: "f \<cdot> a \<noteq> null"
hence "ARR f"
using comp_def null_char by auto
have "rep f \<^bold>\<cdot> rep a \<in> f \<cdot> a"
using assms f Comp_in_comp_ARR comp_def rep_in_ARR null_char by metis
moreover have "rep f \<^bold>\<cdot> rep a \<in> f"
proof -
have "rep f \<in> f"
using \<open>ARR f\<close> rep_in_ARR by auto
moreover have "equiv (rep f \<^bold>\<cdot> rep a) (rep f)"
proof -
have 1: "Par (rep f \<^bold>\<cdot> rep a) (rep f)"
using assms f comp_def mkarr_extensionality Arr_rep_ARR IDE_def null_char
by (metis Cod.simps(4) Dom.simps(4))
moreover have "\<^bold>\<lfloor>rep f \<^bold>\<cdot> rep a\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms f 1 comp_def IDE_def CompDiag_Diag_Ide
Diag_Diagonalize(1) Diag_Diagonalize(2) Diag_Diagonalize(3)
rep_in_ARR
by force
ultimately show ?thesis by auto
qed
ultimately show ?thesis
using \<open>ARR f\<close> ARR_def by auto
qed
ultimately show "f \<cdot> a = f"
using mkarr_memb_ARR comp_def by auto
qed
qed
ultimately show "a \<cdot> a \<noteq> null \<and>
(\<forall>f. (f \<cdot> a \<noteq> null \<longrightarrow> f \<cdot> a = f) \<and> (a \<cdot> f \<noteq> null \<longrightarrow> a \<cdot> f = f))"
by blast
qed
lemma ARR_iff_has_domain:
shows "ARR f \<longleftrightarrow> domains f \<noteq> {}"
proof
assume f: "domains f \<noteq> {}"
show "ARR f" using f domains_def comp_def null_char by auto
next
assume f: "ARR f"
have "Ide (DOM f)"
using f ARR_def by (simp add: Arr_implies_Ide_Dom Arr_rep_ARR)
hence "IDE (mkarr (DOM f))" using IDE_mkarr_Ide by metis
hence "ide (mkarr (DOM f))" using IDE_implies_ide by simp
moreover have "f \<cdot> mkarr (DOM f) = f"
proof -
have 1: "rep f \<^bold>\<cdot> DOM f \<in> f \<cdot> mkarr (DOM f)"
using f Comp_in_comp_ARR
using IDE_implies_ARR Ide_in_Hom rep_in_ARR \<open>IDE (mkarr (DOM f))\<close>
\<open>Ide (DOM f)\<close> Arr_in_mkarr COD_mkarr
by fastforce
moreover have "rep f \<^bold>\<cdot> DOM f \<in> f"
proof -
have 2: "rep f \<in> f" using f rep_in_ARR by simp
moreover have "equiv (rep f \<^bold>\<cdot> DOM f) (rep f)"
by (metis 1 Arr.simps(4) Arr_rep_ARR COD_mkarr Cod.simps(4)
Diagonalize_Comp_Arr_Dom Dom.simps(4) IDE_def Ide_implies_Arr
\<open>IDE (mkarr (DOM f))\<close> \<open>Ide (DOM f)\<close> all_not_in_conv DOM_mkarr comp_def)
ultimately show ?thesis
using f ARR_eqI 1 \<open>ide (mkarr (DOM f))\<close> null_char ide_def by auto
qed
ultimately show ?thesis
using f ARR_eqI \<open>ide (mkarr (DOM f))\<close> null_char ide_def by auto
qed
ultimately show "domains f \<noteq> {}"
using f domains_def not_ARR_empty null_char by auto
qed
lemma ARR_iff_has_codomain:
shows "ARR f \<longleftrightarrow> codomains f \<noteq> {}"
proof
assume f: "codomains f \<noteq> {}"
show "ARR f" using f codomains_def comp_def null_char by auto
next
assume f: "ARR f"
have "Ide (COD f)"
using f ARR_def by (simp add: Arr_rep_ARR Arr_implies_Ide_Cod)
hence "IDE (mkarr (COD f))" using IDE_mkarr_Ide by metis
hence "ide (mkarr (COD f))" using IDE_implies_ide by simp
moreover have "mkarr (COD f) \<cdot> f = f"
proof -
have 1: "COD f \<^bold>\<cdot> rep f \<in> mkarr (COD f) \<cdot> f"
using f Comp_in_comp_ARR
using IDE_implies_ARR Ide_in_Hom rep_in_ARR \<open>IDE (mkarr (COD f))\<close>
\<open>Ide (COD f)\<close> Arr_in_mkarr DOM_mkarr
by fastforce
moreover have "COD f \<^bold>\<cdot> rep f \<in> f"
using 1 null_char norm_rep_ARR norm_memb_eq_rep_ARR mkarr_memb_ARR
\<open>ide (mkarr (COD f))\<close> emptyE equiv_iff_eq_norm mkarr_extensionality ide_def
by metis
ultimately show ?thesis
using f ARR_eqI \<open>ide (mkarr (COD f))\<close> null_char ide_def by auto
qed
ultimately show "codomains f \<noteq> {}"
using codomains_def f not_ARR_empty null_char by auto
qed
lemma arr_iff_ARR:
shows "arr f \<longleftrightarrow> ARR f"
using arr_def ARR_iff_has_domain ARR_iff_has_codomain by simp
text \<open>
The arrows of the category are the equivalence classes of formal arrows.
\<close>
lemma arr_char:
shows "arr f \<longleftrightarrow> f \<noteq> {} \<and> (\<forall>t. t \<in> f \<longrightarrow> f = mkarr t)"
using arr_iff_ARR ARR_def mkarr_def by simp
lemma seq_char:
shows "seq g f \<longleftrightarrow> g \<cdot> f \<noteq> null"
proof
show "g \<cdot> f \<noteq> null \<Longrightarrow> seq g f"
using comp_def null_char Comp_in_comp_ARR rep_in_ARR ARR_mkarr
Arr_rep_ARR arr_iff_ARR
by auto
show "seq g f \<Longrightarrow> g \<cdot> f \<noteq> null"
by auto
qed
lemma seq_char':
shows "seq g f \<longleftrightarrow> ARR f \<and> ARR g \<and> DOM g = COD f"
proof
show "ARR f \<and> ARR g \<and> DOM g = COD f \<Longrightarrow> seq g f"
using comp_def null_char Comp_in_comp_ARR rep_in_ARR ARR_mkarr
Arr_rep_ARR arr_iff_ARR
by auto
have "\<not> (ARR f \<and> ARR g \<and> DOM g = COD f) \<Longrightarrow> g \<cdot> f = null"
using comp_def null_char by auto
thus "seq g f \<Longrightarrow> ARR f \<and> ARR g \<and> DOM g = COD f"
using ext by fastforce
qed
text \<open>
Finally, we can show that the composition \<open>comp\<close> determines a category.
\<close>
interpretation category comp
proof
show "\<And>f. domains f \<noteq> {} \<longleftrightarrow> codomains f \<noteq> {}"
using ARR_iff_has_domain ARR_iff_has_codomain by simp
show 1: "\<And>f g. g \<cdot> f \<noteq> null \<Longrightarrow> seq g f"
using comp_def ARR_comp null_char arr_iff_ARR by metis
fix f g h
show "seq h g \<Longrightarrow> seq (h \<cdot> g) f \<Longrightarrow> seq g f"
using seq_char' by auto
show "seq h (g \<cdot> f) \<Longrightarrow> seq g f \<Longrightarrow> seq h g"
using seq_char' by auto
show "seq g f \<Longrightarrow> seq h g \<Longrightarrow> seq (h \<cdot> g) f"
using seq_char' ARR_comp arr_iff_ARR by auto
show "seq g f \<Longrightarrow> seq h g \<Longrightarrow> (h \<cdot> g) \<cdot> f = h \<cdot> g \<cdot> f"
using seq_char comp_assoc by auto
qed
lemma mkarr_rep [simp]:
assumes "arr f"
shows "mkarr (rep f) = f"
using assms arr_iff_ARR by simp
lemma arr_mkarr [simp]:
assumes "Arr t"
shows "arr (mkarr t)"
using assms by (simp add: ARR_mkarr arr_iff_ARR)
lemma mkarr_memb:
assumes "t \<in> f" and "arr f"
shows "Arr t" and "mkarr t = f"
using assms arr_char mkarr_extensionality by auto
lemma rep_in_arr [simp]:
assumes "arr f"
shows "rep f \<in> f"
using assms by (simp add: rep_in_ARR arr_iff_ARR)
lemma Arr_rep [simp]:
assumes "arr f"
shows "Arr (rep f)"
using assms mkarr_memb rep_in_arr by blast
lemma rep_in_Hom:
assumes "arr f"
shows "rep f \<in> Hom (DOM f) (COD f)"
using assms by simp
lemma norm_memb_eq_rep:
assumes "arr f" and "t \<in> f"
shows "\<^bold>\<parallel>t\<^bold>\<parallel> = rep f"
using assms arr_iff_ARR norm_memb_eq_rep_ARR by auto
lemma norm_rep:
assumes "arr f"
shows "\<^bold>\<parallel>rep f\<^bold>\<parallel> = rep f"
using assms norm_memb_eq_rep by simp
text \<open>
Composition, domain, and codomain on arrows reduce to the corresponding
syntactic operations on their representative terms.
\<close>
lemma comp_mkarr [simp]:
assumes "Arr t" and "Arr u" and "Dom t = Cod u"
shows "mkarr t \<cdot> mkarr u = mkarr (t \<^bold>\<cdot> u)"
using assms
by (metis (no_types, lifting) ARR_mkarr ARR_comp ARR_def Arr_in_mkarr COD_mkarr
Comp_in_comp_ARR DOM_mkarr mkarr_def)
lemma dom_char:
shows "dom f = (if arr f then mkarr (DOM f) else null)"
proof -
have "\<not>arr f \<Longrightarrow> ?thesis"
using dom_def by (simp add: arr_def)
moreover have "arr f \<Longrightarrow> ?thesis"
proof -
assume f: "arr f"
have "dom f = mkarr (DOM f)"
proof (intro dom_eqI)
have 1: "Ide (DOM f)"
using f arr_char by (metis Arr_rep Arr_implies_Ide_Dom)
hence 2: "IDE (mkarr (DOM f))"
using IDE_mkarr_Ide by metis
thus "ide (mkarr (DOM f))" using IDE_implies_ide by simp
moreover show "seq f (mkarr (DOM f))"
proof -
have "f \<cdot> mkarr (DOM f) \<noteq> null"
using f 1 2 ARR_def DOM_mkarr IDE_implies_ARR Ide_in_Hom ARR_comp IDE_def
ARR_iff_has_codomain ARR_iff_has_domain null_char arr_def
by (metis (mono_tags, lifting) mem_Collect_eq)
thus ?thesis using seq_char by simp
qed
qed
thus ?thesis using f by simp
qed
ultimately show ?thesis by blast
qed
lemma dom_simp:
assumes "arr f"
shows "dom f = mkarr (DOM f)"
using assms dom_char by simp
lemma cod_char:
shows "cod f = (if arr f then mkarr (COD f) else null)"
proof -
have "\<not>arr f \<Longrightarrow> ?thesis"
using cod_def by (simp add: arr_def)
moreover have "arr f \<Longrightarrow> ?thesis"
proof -
assume f: "arr f"
have "cod f = mkarr (COD f)"
proof (intro cod_eqI)
have 1: "Ide (COD f)"
using f arr_char by (metis Arr_rep Arr_implies_Ide_Cod)
hence 2: "IDE (mkarr (COD f))"
using IDE_mkarr_Ide by metis
thus "ide (mkarr (COD f))" using IDE_implies_ide by simp
moreover show "seq (mkarr (COD f)) f"
proof -
have "mkarr (COD f) \<cdot> f \<noteq> null"
using f 1 2 ARR_def DOM_mkarr IDE_implies_ARR Ide_in_Hom ARR_comp IDE_def
ARR_iff_has_codomain ARR_iff_has_domain null_char arr_def
by (metis (mono_tags, lifting) mem_Collect_eq)
thus ?thesis using seq_char by simp
qed
qed
thus ?thesis using f by simp
qed
ultimately show ?thesis by blast
qed
lemma cod_simp:
assumes "arr f"
shows "cod f = mkarr (COD f)"
using assms cod_char by simp
lemma Dom_memb:
assumes "arr f" and "t \<in> f"
shows "Dom t = DOM f"
using assms DOM_mkarr mkarr_extensionality arr_char by fastforce
lemma Cod_memb:
assumes "arr f" and "t \<in> f"
shows "Cod t = COD f"
using assms COD_mkarr mkarr_extensionality arr_char by fastforce
lemma dom_mkarr [simp]:
assumes "Arr t"
shows "dom (mkarr t) = mkarr (Dom t)"
using assms dom_char DOM_mkarr arr_mkarr by auto
lemma cod_mkarr [simp]:
assumes "Arr t"
shows "cod (mkarr t) = mkarr (Cod t)"
using assms cod_char COD_mkarr arr_mkarr by auto
lemma mkarr_in_hom:
assumes "Arr t"
shows "\<guillemotleft>mkarr t : mkarr (Dom t) \<rightarrow> mkarr (Cod t)\<guillemotright>"
using assms arr_mkarr dom_mkarr cod_mkarr by auto
lemma DOM_in_dom [intro]:
assumes "arr f"
shows "DOM f \<in> dom f"
using assms dom_char
by (metis Arr_in_mkarr mkarr_extensionality ideD(1) ide_dom not_arr_null null_char)
lemma COD_in_cod [intro]:
assumes "arr f"
shows "COD f \<in> cod f"
using assms cod_char
by (metis Arr_in_mkarr mkarr_extensionality ideD(1) ide_cod not_arr_null null_char)
lemma DOM_dom:
assumes "arr f"
shows "DOM (dom f) = DOM f"
using assms Arr_rep Arr_implies_Ide_Dom Ide_implies_Arr dom_char rep_mkarr Par_Arr_norm
Ide_in_Hom
by simp
lemma DOM_cod:
assumes "arr f"
shows "DOM (cod f) = COD f"
using assms Arr_rep Arr_implies_Ide_Cod Ide_implies_Arr cod_char rep_mkarr Par_Arr_norm
Ide_in_Hom
by simp
lemma memb_equiv:
assumes "arr f" and "t \<in> f" and "u \<in> f"
shows "Par t u" and "\<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
proof -
show "Par t u"
using assms Cod_memb Dom_memb mkarr_memb(1) by metis
show "\<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
using assms arr_iff_ARR ARR_def by auto
qed
text \<open>
Two arrows can be proved equal by showing that they are parallel and
have representatives with identical diagonalizations.
\<close>
lemma arr_eqI:
assumes "par f g" and "t \<in> f" and "u \<in> g" and "\<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
shows "f = g"
proof -
have "Arr t \<and> Arr u" using assms mkarr_memb(1) by blast
moreover have "Dom t = Dom u \<and> Cod t = Cod u"
using assms Dom_memb Cod_memb comp_def arr_char comp_arr_dom comp_cod_arr
by (metis (full_types))
ultimately have "Par t u" by simp
thus ?thesis
using assms arr_char by (metis rep_mkarr rep_in_arr equiv_iff_eq_norm)
qed
lemma comp_char:
shows "f \<cdot> g = (if seq f g then mkarr (rep f \<^bold>\<cdot> rep g) else null)"
using comp_def seq_char arr_char by meson
text \<open>
The mapping that takes identity terms to their equivalence classes is injective.
\<close>
lemma mkarr_inj_on_Ide:
assumes "Ide t" and "Ide u" and "mkarr t = mkarr u"
shows "t = u"
using assms
by (metis (mono_tags, lifting) COD_mkarr Ide_in_Hom mem_Collect_eq)
lemma Comp_in_comp [intro]:
assumes "arr f" and "g \<in> hom (dom g) (dom f)" and "t \<in> f" and "u \<in> g"
shows "t \<^bold>\<cdot> u \<in> f \<cdot> g"
proof -
have "ARR f" using assms arr_iff_ARR by simp
moreover have "ARR g" using assms arr_iff_ARR by auto
moreover have "DOM f = COD g"
using assms dom_char cod_char mkarr_inj_on_Ide Arr_implies_Ide_Cod Arr_implies_Ide_Dom
by force
ultimately show ?thesis using assms Comp_in_comp_ARR by simp
qed
text \<open>
An arrow is defined to be ``canonical'' if some (equivalently, all) its representatives
diagonalize to an identity term.
\<close>
definition can
where "can f \<equiv> arr f \<and> (\<exists>t. t \<in> f \<and> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>)"
lemma can_def_alt:
shows "can f \<longleftrightarrow> arr f \<and> (\<forall>t. t \<in> f \<longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>)"
proof
assume "arr f \<and> (\<forall>t. t \<in> f \<longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>)"
thus "can f" using can_def arr_char by fastforce
next
assume f: "can f"
show "arr f \<and> (\<forall>t. t \<in> f \<longrightarrow> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>)"
proof -
obtain t where t: "t \<in> f \<and> Ide \<^bold>\<lfloor>t\<^bold>\<rfloor>" using f can_def by auto
have "ARR f" using f can_def arr_char ARR_def mkarr_def by simp
hence "\<forall>u. u \<in> f \<longrightarrow> \<^bold>\<parallel>u\<^bold>\<parallel> = \<^bold>\<parallel>t\<^bold>\<parallel>" using t unique_norm by auto
hence "\<forall>u. u \<in> f \<longrightarrow> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>u\<^bold>\<rfloor>"
using t by (metis \<open>ARR f\<close> equiv_iff_eq_norm arr_iff_ARR mkarr_memb(1))
hence "\<forall>u. u \<in> f \<longrightarrow> Ide \<^bold>\<lfloor>u\<^bold>\<rfloor>"
using t by metis
thus ?thesis using f can_def by blast
qed
qed
lemma can_implies_arr:
assumes "can f"
shows "arr f"
using assms can_def by auto
text \<open>
The identities of the category are precisely the canonical endo-arrows.
\<close>
lemma ide_char:
shows "ide f \<longleftrightarrow> can f \<and> dom f = cod f"
proof
assume f: "ide f"
show "can f \<and> dom f = cod f"
using f can_def arr_char dom_char cod_char IDE_def Arr_implies_Ide_Cod can_def_alt
Arr_rep IDE_mkarr_Ide
by (metis ideD(1) ideD(3))
next
assume f: "can f \<and> dom f = cod f"
show "ide f"
proof -
have "f = dom f"
proof (intro arr_eqI)
show "par f (dom f)" using f can_def by simp
show "rep f \<in> f" using f can_def by simp
show "DOM f \<in> dom f" using f can_def by auto
show "\<^bold>\<lfloor>rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>DOM f\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<in> Hom \<^bold>\<lfloor>DOM f\<^bold>\<rfloor> \<^bold>\<lfloor>COD f\<^bold>\<rfloor>"
using f can_def Diagonalize_in_Hom by simp
moreover have "Ide \<^bold>\<lfloor>rep f\<^bold>\<rfloor>" using f can_def_alt rep_in_arr by simp
ultimately show ?thesis
using f can_def Ide_in_Hom by simp
qed
qed
thus ?thesis using f can_implies_arr ide_dom [of f] by auto
qed
qed
lemma ide_iff_IDE:
shows "ide a \<longleftrightarrow> IDE a"
using ide_char IDE_def can_def_alt arr_iff_ARR dom_char cod_char mkarr_inj_on_Ide
Arr_implies_Ide_Cod Arr_implies_Ide_Dom Arr_rep
by auto
lemma ide_mkarr_Ide:
assumes "Ide a"
shows "ide (mkarr a)"
using assms IDE_mkarr_Ide ide_iff_IDE by simp
lemma rep_dom:
assumes "arr f"
shows "rep (dom f) = \<^bold>\<parallel>DOM f\<^bold>\<parallel>"
using assms dom_simp rep_mkarr Arr_rep Arr_implies_Ide_Dom by simp
lemma rep_cod:
assumes "arr f"
shows "rep (cod f) = \<^bold>\<parallel>COD f\<^bold>\<parallel>"
using assms cod_simp rep_mkarr Arr_rep Arr_implies_Ide_Cod by simp
lemma rep_preserves_seq:
assumes "seq g f"
shows "Seq (rep g) (rep f)"
using assms Arr_rep dom_char cod_char mkarr_inj_on_Ide Arr_implies_Ide_Dom
Arr_implies_Ide_Cod
by auto
lemma rep_comp:
assumes "seq g f"
shows "rep (g \<cdot> f) = \<^bold>\<parallel>rep g \<^bold>\<cdot> rep f\<^bold>\<parallel>"
proof -
have "rep (g \<cdot> f) = rep (mkarr (rep g \<^bold>\<cdot> rep f))"
using assms comp_char by metis
also have "... = \<^bold>\<parallel>rep g \<^bold>\<cdot> rep f\<^bold>\<parallel>"
using assms rep_preserves_seq rep_mkarr by simp
finally show ?thesis by blast
qed
text \<open>
The equivalence classes of canonical terms are canonical arrows.
\<close>
lemma can_mkarr_Can:
assumes "Can t"
shows "can (mkarr t)"
using assms Arr_in_mkarr Can_implies_Arr Ide_Diagonalize_Can arr_mkarr can_def by blast
lemma ide_implies_can:
assumes "ide a"
shows "can a"
using assms ide_char by blast
lemma Can_rep_can:
assumes "can f"
shows "Can (rep f)"
proof -
have "Can \<^bold>\<parallel>rep f\<^bold>\<parallel>"
using assms can_def_alt Can_norm_iff_Ide_Diagonalize by auto
moreover have "rep f = \<^bold>\<parallel>rep f\<^bold>\<parallel>"
using assms can_implies_arr norm_rep by simp
ultimately show ?thesis by simp
qed
text \<open>
Parallel canonical arrows are identical.
\<close>
lemma can_coherence:
assumes "par f g" and "can f" and "can g"
shows "f = g"
proof -
have "\<^bold>\<lfloor>rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>rep g\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>DOM f\<^bold>\<rfloor>"
using assms Ide_Diagonalize_Can Can_rep_can Diagonalize_in_Hom Ide_in_Hom by force
also have "... = \<^bold>\<lfloor>DOM g\<^bold>\<rfloor>"
using assms dom_char equiv_iff_eq_norm
by (metis DOM_in_dom mkarr_memb(1) rep_mkarr arr_dom_iff_arr)
also have "... = \<^bold>\<lfloor>rep g\<^bold>\<rfloor>"
using assms Ide_Diagonalize_Can Can_rep_can Diagonalize_in_Hom Ide_in_Hom by force
finally show ?thesis by blast
qed
hence "rep f = rep g"
using assms rep_in_arr norm_memb_eq_rep equiv_iff_eq_norm
by (metis (no_types, lifting) arr_eqI)
thus ?thesis
using assms arr_eqI [of f g] rep_in_arr [of f] rep_in_arr [of g] by metis
qed
text \<open>
Canonical arrows are invertible, and their inverses can be obtained syntactically.
\<close>
lemma inverse_arrows_can:
assumes "can f"
shows "inverse_arrows f (mkarr (Inv (DOM f\<^bold>\<down>) \<^bold>\<cdot> \<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<cdot> COD f\<^bold>\<down>))"
proof
let ?t = "(Inv (DOM f\<^bold>\<down>) \<^bold>\<cdot> \<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<cdot> COD f\<^bold>\<down>)"
have 1: "rep f \<in> f \<and> Arr (rep f) \<and> Can (rep f) \<and> Ide \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms can_def_alt rep_in_arr rep_in_arr(1) Can_rep_can by simp
hence 2: "\<^bold>\<lfloor>DOM f\<^bold>\<rfloor> = \<^bold>\<lfloor>COD f\<^bold>\<rfloor>"
using Diagonalize_in_Hom [of "rep f"] Ide_in_Hom by auto
have 3: "Can ?t"
using assms 1 2 Can_red Ide_implies_Can Diagonalize_in_Hom Inv_preserves_Can
Arr_implies_Ide_Cod Arr_implies_Ide_Dom Diag_Diagonalize
by simp
have 4: "DOM f = Cod ?t"
using assms can_def Can_red
by (simp add: Arr_implies_Ide_Dom Inv_preserves_Can(3))
have 5: "COD f = Dom ?t"
using assms can_def Can_red Arr_rep Arr_implies_Ide_Cod by simp
have 6: "antipar f (mkarr ?t)"
using assms 3 4 5 dom_char cod_char can_def cod_mkarr dom_mkarr Can_implies_Arr
by simp
show "ide (f \<cdot> mkarr ?t)"
proof -
have 7: "par (f \<cdot> mkarr ?t) (dom (f \<cdot> mkarr ?t))"
using assms 6 by auto
moreover have "can (f \<cdot> mkarr ?t)"
proof -
have 8: "Comp (rep f) ?t \<in> (f \<cdot> mkarr ?t)"
using assms 1 3 4 6 can_implies_arr Arr_in_mkarr COD_mkarr Comp_in_comp_ARR
Can_implies_Arr arr_iff_ARR seq_char'
by meson
moreover have "Can (rep f \<^bold>\<cdot> ?t)"
using 1 3 7 8 mkarr_memb(1) by (metis Arr.simps(4) Can.simps(4))
ultimately show ?thesis
using can_mkarr_Can 7 mkarr_memb(2) by metis
qed
moreover have "can (dom (f \<cdot> mkarr ?t))"
using 7 ide_implies_can by force
ultimately have "f \<cdot> mkarr ?t = dom (f \<cdot> mkarr ?t)"
using can_coherence by meson
thus ?thesis
using 7 ide_dom by metis
qed
show "ide (mkarr ?t \<cdot> f)"
proof -
have 7: "par (mkarr ?t \<cdot> f) (cod (mkarr ?t \<cdot> f))"
using assms 6 by auto
moreover have "can (mkarr ?t \<cdot> f)"
proof -
have 8: "Comp ?t (rep f) \<in> mkarr ?t \<cdot> f"
using assms 1 3 6 7 Arr_in_mkarr Comp_in_comp_ARR Can_implies_Arr arr_char
comp_def
by meson
moreover have "Can (?t \<^bold>\<cdot> rep f)"
using 1 3 7 8 mkarr_memb(1) by (metis Arr.simps(4) Can.simps(4))
ultimately show ?thesis
using can_mkarr_Can 7 mkarr_memb(2) by metis
qed
moreover have "can (cod (mkarr ?t \<cdot> f))"
using 7 ide_implies_can by force
ultimately have "mkarr ?t \<cdot> f = cod (mkarr ?t \<cdot> f)"
using can_coherence by meson
thus ?thesis
using 7 can_implies_arr ide_cod by metis
qed
qed
lemma inv_mkarr [simp]:
assumes "Can t"
shows "inv (mkarr t) = mkarr (Inv t)"
proof -
have t: "Can t \<and> Arr t \<and> Can (Inv t) \<and> Arr (Inv t) \<and> Ide (Dom t) \<and> Ide (Cod t)"
using assms Can_implies_Arr Arr_implies_Ide_Dom Arr_implies_Ide_Cod
Inv_preserves_Can
by simp
have "inverse_arrows (mkarr t) (mkarr (Inv t))"
proof
show "ide (mkarr t \<cdot> mkarr (Inv t))"
proof -
have "mkarr (Cod t) = mkarr (Comp t (Inv t))"
using t Inv_in_Hom Ide_in_Hom Diagonalize_Inv Diag_Diagonalize Diagonalize_preserves_Can
by (intro mkarr_eqI, auto)
also have "... = mkarr t \<cdot> mkarr (Inv t)"
using t comp_mkarr Inv_in_Hom by simp
finally have "mkarr (Cod t) = mkarr t \<cdot> mkarr (Inv t)"
by blast
thus ?thesis using t ide_mkarr_Ide [of "Cod t"] by simp
qed
show "ide (mkarr (Inv t) \<cdot> mkarr t)"
proof -
have "mkarr (Dom t) = mkarr (Inv t \<^bold>\<cdot> t)"
using t Inv_in_Hom Ide_in_Hom Diagonalize_Inv Diag_Diagonalize Diagonalize_preserves_Can
by (intro mkarr_eqI, auto)
also have "... = mkarr (Inv t) \<cdot> mkarr t"
using t comp_mkarr Inv_in_Hom by simp
finally have "mkarr (Dom t) = mkarr (Inv t) \<cdot> mkarr t"
by blast
thus ?thesis using t ide_mkarr_Ide [of "Dom t"] by simp
qed
qed
thus ?thesis using inverse_unique by auto
qed
lemma iso_can:
assumes "can f"
shows "iso f"
using assms inverse_arrows_can by auto
text \<open>
The following function produces the unique canonical arrow between two given objects,
if such an arrow exists.
\<close>
definition mkcan
where "mkcan a b = mkarr (Inv (COD b\<^bold>\<down>) \<^bold>\<cdot> (DOM a\<^bold>\<down>))"
lemma can_mkcan:
assumes "ide a" and "ide b" and "\<^bold>\<lfloor>DOM a\<^bold>\<rfloor> = \<^bold>\<lfloor>COD b\<^bold>\<rfloor>"
shows "can (mkcan a b)" and "\<guillemotleft>mkcan a b : a \<rightarrow> b\<guillemotright>"
proof -
show "can (mkcan a b)"
using assms mkcan_def Arr_rep Arr_implies_Ide_Dom Arr_implies_Ide_Cod Can_red
Inv_preserves_Can can_mkarr_Can
by simp
show "\<guillemotleft>mkcan a b : a \<rightarrow> b\<guillemotright>"
using assms mkcan_def Arr_rep Arr_implies_Ide_Dom Arr_implies_Ide_Cod Can_red Inv_in_Hom
dom_char [of a] cod_char [of b] mkarr_rep mkarr_in_hom can_implies_arr
by auto
qed
lemma dom_mkcan:
assumes "ide a" and "ide b" and "\<^bold>\<lfloor>DOM a\<^bold>\<rfloor> = \<^bold>\<lfloor>COD b\<^bold>\<rfloor>"
shows "dom (mkcan a b) = a"
using assms can_mkcan by blast
lemma cod_mkcan:
assumes "ide a" and "ide b" and "\<^bold>\<lfloor>DOM a\<^bold>\<rfloor> = \<^bold>\<lfloor>COD b\<^bold>\<rfloor>"
shows "cod (mkcan a b) = b"
using assms can_mkcan by blast
lemma can_coherence':
assumes "can f"
shows "mkcan (dom f) (cod f) = f"
proof -
have "Ide \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms Ide_Diagonalize_Can Can_rep_can by simp
hence "Dom \<^bold>\<lfloor>rep f\<^bold>\<rfloor> = Cod \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using Ide_in_Hom by simp
hence "\<^bold>\<lfloor>DOM f\<^bold>\<rfloor> = \<^bold>\<lfloor>COD f\<^bold>\<rfloor>"
using assms can_implies_arr Arr_rep Diagonalize_in_Hom by simp
moreover have "DOM f = DOM (dom f)"
using assms can_implies_arr dom_char rep_mkarr Arr_implies_Ide_Dom Ide_implies_Arr
Par_Arr_norm [of "DOM f"] Ide_in_Hom
by auto
moreover have "COD f = COD (cod f)"
using assms can_implies_arr cod_char rep_mkarr Arr_implies_Ide_Cod Ide_implies_Arr
Par_Arr_norm [of "COD f"] Ide_in_Hom
by auto
ultimately have "can (mkcan (dom f) (cod f)) \<and> par f (mkcan (dom f) (cod f))"
using assms can_implies_arr can_mkcan dom_mkcan cod_mkcan by simp
thus ?thesis using assms can_coherence by blast
qed
lemma Ide_Diagonalize_rep_ide:
assumes "ide a"
shows "Ide \<^bold>\<lfloor>rep a\<^bold>\<rfloor>"
using assms ide_implies_can can_def_alt rep_in_arr by simp
lemma Diagonalize_DOM:
assumes "arr f"
shows "\<^bold>\<lfloor>DOM f\<^bold>\<rfloor> = Dom \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms Diag_Diagonalize by simp
lemma Diagonalize_COD:
assumes "arr f"
shows "\<^bold>\<lfloor>COD f\<^bold>\<rfloor> = Cod \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms Diag_Diagonalize by simp
lemma Diagonalize_rep_preserves_seq:
assumes "seq g f"
shows "Seq \<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms Diagonalize_DOM Diagonalize_COD Diag_implies_Arr Diag_Diagonalize(1)
rep_preserves_seq
by force
lemma Dom_Diagonalize_rep:
assumes "arr f"
shows "Dom \<^bold>\<lfloor>rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>rep (dom f)\<^bold>\<rfloor>"
using assms Diagonalize_rep_preserves_seq [of f "dom f"] Ide_Diagonalize_rep_ide Ide_in_Hom
by simp
lemma Cod_Diagonalize_rep:
assumes "arr f"
shows "Cod \<^bold>\<lfloor>rep f\<^bold>\<rfloor> = \<^bold>\<lfloor>rep (cod f)\<^bold>\<rfloor>"
using assms Diagonalize_rep_preserves_seq [of "cod f" f] Ide_Diagonalize_rep_ide Ide_in_Hom
by simp
lemma mkarr_Diagonalize_rep:
assumes "arr f" and "Diag (DOM f)" and "Diag (COD f)"
shows "mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor> = f"
proof -
have "mkarr (rep f) = mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms rep_in_Hom Diagonalize_in_Hom Diag_Diagonalize Diagonalize_Diag
by (intro mkarr_eqI, simp_all)
thus ?thesis using assms mkarr_rep by auto
qed
text \<open>
We define tensor product of arrows via the constructor @{term Tensor} on terms.
\<close>
definition tensor\<^sub>F\<^sub>M\<^sub>C (infixr "\<otimes>" 53)
where "f \<otimes> g \<equiv> (if arr f \<and> arr g then mkarr (rep f \<^bold>\<otimes> rep g) else null)"
lemma arr_tensor [simp]:
assumes "arr f" and "arr g"
shows "arr (f \<otimes> g)"
using assms tensor\<^sub>F\<^sub>M\<^sub>C_def arr_mkarr by simp
lemma rep_tensor:
assumes "arr f" and "arr g"
shows "rep (f \<otimes> g) = \<^bold>\<parallel>rep f \<^bold>\<otimes> rep g\<^bold>\<parallel>"
using assms tensor\<^sub>F\<^sub>M\<^sub>C_def rep_mkarr by simp
lemma Par_memb_rep:
assumes "arr f" and "t \<in> f"
shows "Par t (rep f)"
using assms mkarr_memb apply simp
using rep_in_Hom Dom_memb Cod_memb by metis
lemma Tensor_in_tensor [intro]:
assumes "arr f" and "arr g" and "t \<in> f" and "u \<in> g"
shows "t \<^bold>\<otimes> u \<in> f \<otimes> g"
proof -
have "equiv (t \<^bold>\<otimes> u) (rep f \<^bold>\<otimes> rep g)"
proof -
have 1: "Par (t \<^bold>\<otimes> u) (rep f \<^bold>\<otimes> rep g)"
proof -
have "Par t (rep f) \<and> Par u (rep g)" using assms Par_memb_rep by blast
thus ?thesis by simp
qed
moreover have "\<^bold>\<lfloor>t \<^bold>\<otimes> u\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f \<^bold>\<otimes> rep g\<^bold>\<rfloor>"
using assms 1 equiv_iff_eq_norm rep_mkarr norm_norm mkarr_memb(2)
by (metis Arr.simps(3) Diagonalize.simps(3))
ultimately show ?thesis by simp
qed
thus ?thesis
using assms tensor\<^sub>F\<^sub>M\<^sub>C_def mkarr_def by simp
qed
lemma DOM_tensor [simp]:
assumes "arr f" and "arr g"
shows "DOM (f \<otimes> g) = DOM f \<^bold>\<otimes> DOM g"
by (metis (no_types, lifting) DOM_mkarr Dom.simps(3) mkarr_extensionality arr_char
arr_tensor assms(1) assms(2) tensor\<^sub>F\<^sub>M\<^sub>C_def)
lemma COD_tensor [simp]:
assumes "arr f" and "arr g"
shows "COD (f \<otimes> g) = COD f \<^bold>\<otimes> COD g"
by (metis (no_types, lifting) COD_mkarr Cod.simps(3) mkarr_extensionality arr_char
arr_tensor assms(1) assms(2) tensor\<^sub>F\<^sub>M\<^sub>C_def)
lemma tensor_in_hom [simp]:
assumes "\<guillemotleft>f : a \<rightarrow> b\<guillemotright>" and "\<guillemotleft>g : c \<rightarrow> d\<guillemotright>"
shows "\<guillemotleft>f \<otimes> g : a \<otimes> c \<rightarrow> b \<otimes> d\<guillemotright>"
proof -
have f: "arr f \<and> dom f = a \<and> cod f = b" using assms(1) by auto
have g: "arr g \<and> dom g = c \<and> cod g = d" using assms(2) by auto
have "dom (f \<otimes> g) = dom f \<otimes> dom g"
using f g arr_tensor dom_char Tensor_in_tensor [of "dom f" "dom g" "DOM f" "DOM g"]
DOM_in_dom mkarr_memb(2) DOM_tensor arr_dom_iff_arr
by metis
moreover have "cod (f \<otimes> g) = cod f \<otimes> cod g"
using f g arr_tensor cod_char Tensor_in_tensor [of "cod f" "cod g" "COD f" "COD g"]
COD_in_cod mkarr_memb(2) COD_tensor arr_cod_iff_arr
by metis
ultimately show ?thesis using assms arr_tensor by blast
qed
lemma dom_tensor [simp]:
assumes "arr f" and "arr g"
shows "dom (f \<otimes> g) = dom f \<otimes> dom g"
using assms tensor_in_hom [of f] by blast
lemma cod_tensor [simp]:
assumes "arr f" and "arr g"
shows "cod (f \<otimes> g) = cod f \<otimes> cod g"
using assms tensor_in_hom [of f] by blast
lemma tensor_mkarr [simp]:
assumes "Arr t" and "Arr u"
shows "mkarr t \<otimes> mkarr u = mkarr (t \<^bold>\<otimes> u)"
using assms by (meson Tensor_in_tensor arr_char Arr_in_mkarr arr_mkarr arr_tensor)
lemma tensor_preserves_ide:
assumes "ide a" and "ide b"
shows "ide (a \<otimes> b)"
proof -
have "can (a \<otimes> b)"
using assms tensor\<^sub>F\<^sub>M\<^sub>C_def Can_rep_can ide_implies_can can_mkarr_Can by simp
moreover have "dom (a \<otimes> b) = cod (a \<otimes> b)"
using assms tensor_in_hom by simp
ultimately show ?thesis using ide_char by metis
qed
lemma tensor_preserves_can:
assumes "can f" and "can g"
shows "can (f \<otimes> g)"
using assms can_implies_arr Can_rep_can tensor\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can by simp
lemma comp_preserves_can:
assumes "can f" and "can g" and "dom f = cod g"
shows "can (f \<cdot> g)"
proof -
have 1: "ARR f \<and> ARR g \<and> DOM f = COD g"
using assms can_implies_arr arr_iff_ARR Arr_implies_Ide_Cod Arr_implies_Ide_Dom
mkarr_inj_on_Ide cod_char dom_char
by simp
hence "Can (rep f \<^bold>\<cdot> rep g)"
using assms can_implies_arr Can_rep_can by force
thus ?thesis
using assms 1 can_implies_arr comp_char can_mkarr_Can seq_char' by simp
qed
text \<open>
The remaining structure required of a monoidal category is also defined syntactically.
\<close>
definition unity\<^sub>F\<^sub>M\<^sub>C :: "'c arr" ("\<I>")
where "\<I> = mkarr \<^bold>\<I>"
definition lunit\<^sub>F\<^sub>M\<^sub>C :: "'c arr \<Rightarrow> 'c arr" ("\<l>[_]")
where "\<l>[a] = mkarr \<^bold>\<l>\<^bold>[rep a\<^bold>]"
definition runit\<^sub>F\<^sub>M\<^sub>C :: "'c arr \<Rightarrow> 'c arr" ("\<r>[_]")
where "\<r>[a] = mkarr \<^bold>\<r>\<^bold>[rep a\<^bold>]"
definition assoc\<^sub>F\<^sub>M\<^sub>C :: "'c arr \<Rightarrow> 'c arr \<Rightarrow> 'c arr \<Rightarrow> 'c arr" ("\<a>[_, _, _]")
where "\<a>[a, b, c] = mkarr \<^bold>\<a>\<^bold>[rep a, rep b, rep c\<^bold>]"
lemma can_lunit:
assumes "ide a"
shows "can \<l>[a]"
using assms lunit\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can
by (simp add: Can_rep_can ide_implies_can)
lemma lunit_in_hom:
assumes "ide a"
shows "\<guillemotleft>\<l>[a] : \<I> \<otimes> a \<rightarrow> a\<guillemotright>"
proof -
have "dom \<l>[a] = \<I> \<otimes> a"
using assms lunit\<^sub>F\<^sub>M\<^sub>C_def unity\<^sub>F\<^sub>M\<^sub>C_def Ide_implies_Arr dom_mkarr dom_char tensor_mkarr
Arr_rep
by (metis Arr.simps(2) Arr.simps(5) Arr_implies_Ide_Dom Dom.simps(5)
ideD(1) ideD(2))
moreover have "cod \<l>[a] = a"
using assms lunit\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) cod_mkarr cod_char ideD(3) by auto
ultimately show ?thesis
using assms arr_cod_iff_arr by (intro in_homI, fastforce)
qed
lemma arr_lunit [simp]:
assumes "ide a"
shows "arr \<l>[a]"
using assms can_lunit can_implies_arr by simp
lemma dom_lunit [simp]:
assumes "ide a"
shows "dom \<l>[a] = \<I> \<otimes> a"
using assms lunit_in_hom by auto
lemma cod_lunit [simp]:
assumes "ide a"
shows "cod \<l>[a] = a"
using assms lunit_in_hom by auto
lemma can_runit:
assumes "ide a"
shows "can \<r>[a]"
using assms runit\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can
by (simp add: Can_rep_can ide_implies_can)
lemma runit_in_hom [simp]:
assumes "ide a"
shows "\<guillemotleft>\<r>[a] : a \<otimes> \<I> \<rightarrow> a\<guillemotright>"
proof -
have "dom \<r>[a] = a \<otimes> \<I>"
using assms Arr_rep Arr.simps(2) Arr.simps(7) Arr_implies_Ide_Dom Dom.simps(7)
Ide_implies_Arr dom_mkarr dom_char ideD(1) ideD(2) runit\<^sub>F\<^sub>M\<^sub>C_def tensor_mkarr
unity\<^sub>F\<^sub>M\<^sub>C_def
by metis
moreover have "cod \<r>[a] = a"
using assms runit\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) cod_mkarr cod_char ideD(3) by auto
ultimately show ?thesis
using assms arr_cod_iff_arr by (intro in_homI, fastforce)
qed
lemma arr_runit [simp]:
assumes "ide a"
shows "arr \<r>[a]"
using assms can_runit can_implies_arr by simp
lemma dom_runit [simp]:
assumes "ide a"
shows "dom \<r>[a] = a \<otimes> \<I>"
using assms runit_in_hom by blast
lemma cod_runit [simp]:
assumes "ide a"
shows "cod \<r>[a] = a"
using assms runit_in_hom by blast
lemma can_assoc:
assumes "ide a" and "ide b" and "ide c"
shows "can \<a>[a, b, c]"
using assms assoc\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can
by (simp add: Can_rep_can ide_implies_can)
lemma assoc_in_hom:
assumes "ide a" and "ide b" and "ide c"
shows "\<guillemotleft>\<a>[a, b, c] : (a \<otimes> b) \<otimes> c \<rightarrow> a \<otimes> b \<otimes> c\<guillemotright>"
proof -
have "dom \<a>[a, b, c] = (a \<otimes> b) \<otimes> c"
proof -
have "dom \<a>[a, b, c] = mkarr (Dom \<^bold>\<a>\<^bold>[rep a, rep b, rep c\<^bold>])"
using assms assoc\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) by simp
also have "... = mkarr ((DOM a \<^bold>\<otimes> DOM b) \<^bold>\<otimes> DOM c)"
by simp
also have "... = (a \<otimes> b) \<otimes> c"
by (metis mkarr_extensionality arr_tensor assms dom_char
ideD(1) ideD(2) not_arr_null null_char tensor_mkarr)
finally show ?thesis by blast
qed
moreover have "cod \<a>[a, b, c] = a \<otimes> b \<otimes> c"
proof -
have "cod \<a>[a, b, c] = mkarr (Cod \<^bold>\<a>\<^bold>[rep a, rep b, rep c\<^bold>])"
using assms assoc\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) by simp
also have "... = mkarr (COD a \<^bold>\<otimes> COD b \<^bold>\<otimes> COD c)"
by simp
also have "... = a \<otimes> b \<otimes> c"
by (metis mkarr_extensionality arr_tensor assms(1) assms(2) assms(3) cod_char
ideD(1) ideD(3) not_arr_null null_char tensor_mkarr)
finally show ?thesis by blast
qed
moreover have "arr \<a>[a, b, c]"
using assms assoc\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) arr_mkarr by simp
ultimately show ?thesis by auto
qed
lemma arr_assoc [simp]:
assumes "ide a" and "ide b" and "ide c"
shows "arr \<a>[a, b, c]"
using assms can_assoc can_implies_arr by simp
lemma dom_assoc [simp]:
assumes "ide a" and "ide b" and "ide c"
shows "dom \<a>[a, b, c] = (a \<otimes> b) \<otimes> c"
using assms assoc_in_hom by blast
lemma cod_assoc [simp]:
assumes "ide a" and "ide b" and "ide c"
shows "cod \<a>[a, b, c] = a \<otimes> b \<otimes> c"
using assms assoc_in_hom by blast
lemma ide_unity [simp]:
shows "ide \<I>"
using unity\<^sub>F\<^sub>M\<^sub>C_def Arr.simps(2) Dom.simps(2) arr_mkarr dom_mkarr ide_dom
by metis
lemma Unity_in_unity [simp]:
shows "\<^bold>\<I> \<in> \<I>"
using unity\<^sub>F\<^sub>M\<^sub>C_def Arr_in_mkarr by simp
lemma rep_unity [simp]:
shows "rep \<I> = \<^bold>\<parallel>\<^bold>\<I>\<^bold>\<parallel>"
using unity\<^sub>F\<^sub>M\<^sub>C_def rep_mkarr by simp
lemma Lunit_in_lunit [intro]:
assumes "arr f" and "t \<in> f"
shows "\<^bold>\<l>\<^bold>[t\<^bold>] \<in> \<l>[f]"
proof -
have "Arr t \<and> Arr (rep f) \<and> Dom t = DOM f \<and> Cod t = COD f \<and> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms
by (metis mkarr_memb(1) mkarr_memb(2) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm
norm_rep)
thus ?thesis
by (simp add: mkarr_def lunit\<^sub>F\<^sub>M\<^sub>C_def)
qed
lemma Runit_in_runit [intro]:
assumes "arr f" and "t \<in> f"
shows "\<^bold>\<r>\<^bold>[t\<^bold>] \<in> \<r>[f]"
proof -
have "Arr t \<and> Arr (rep f) \<and> Dom t = DOM f \<and> Cod t = COD f \<and> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms
by (metis mkarr_memb(1) mkarr_memb(2) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm
norm_rep)
thus ?thesis
by (simp add: mkarr_def runit\<^sub>F\<^sub>M\<^sub>C_def)
qed
lemma Assoc_in_assoc [intro]:
assumes "arr f" and "arr g" and "arr h"
and "t \<in> f" and "u \<in> g" and "v \<in> h"
shows "\<^bold>\<a>\<^bold>[t, u, v\<^bold>] \<in> \<a>[f, g, h]"
proof -
have "Arr t \<and> Arr (rep f) \<and> Dom t = DOM f \<and> Cod t = COD f \<and> \<^bold>\<lfloor>t\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using assms
by (metis mkarr_memb(1) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm mkarr_memb(2)
norm_rep)
moreover have "Arr u \<and> Arr (rep g) \<and> Dom u = DOM g \<and> Cod u = COD g \<and>
\<^bold>\<lfloor>u\<^bold>\<rfloor> = \<^bold>\<lfloor>rep g\<^bold>\<rfloor>"
using assms
by (metis mkarr_memb(1) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm mkarr_memb(2)
norm_rep)
moreover have "Arr v \<and> Arr (rep h) \<and> Dom v = DOM h \<and> Cod v = COD h \<and>
\<^bold>\<lfloor>v\<^bold>\<rfloor> = \<^bold>\<lfloor>rep h\<^bold>\<rfloor>"
using assms
by (metis mkarr_memb(1) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm mkarr_memb(2)
norm_rep)
ultimately show ?thesis
using assoc\<^sub>F\<^sub>M\<^sub>C_def mkarr_def by simp
qed
text \<open>
At last, we can show that we've constructed a monoidal category.
\<close>
interpretation EMC: elementary_monoidal_category
comp tensor\<^sub>F\<^sub>M\<^sub>C unity\<^sub>F\<^sub>M\<^sub>C lunit\<^sub>F\<^sub>M\<^sub>C runit\<^sub>F\<^sub>M\<^sub>C assoc\<^sub>F\<^sub>M\<^sub>C
proof
show "ide \<I>" using ide_unity by auto
show "\<And>a. ide a \<Longrightarrow> \<guillemotleft>\<l>[a] : \<I> \<otimes> a \<rightarrow> a\<guillemotright>" by auto
show "\<And>a. ide a \<Longrightarrow> \<guillemotleft>\<r>[a] : a \<otimes> \<I> \<rightarrow> a\<guillemotright>" by auto
show "\<And>a. ide a \<Longrightarrow> iso \<l>[a]" using can_lunit iso_can by auto
show "\<And>a. ide a \<Longrightarrow> iso \<r>[a]" using can_runit iso_can by auto
show "\<And>a b c. \<lbrakk> ide a; ide b; ide c \<rbrakk> \<Longrightarrow> \<guillemotleft>\<a>[a, b, c] : (a \<otimes> b) \<otimes> c \<rightarrow> a \<otimes> b \<otimes> c\<guillemotright>" by auto
show "\<And>a b c. \<lbrakk> ide a; ide b; ide c \<rbrakk> \<Longrightarrow> iso \<a>[a, b, c]" using can_assoc iso_can by auto
show "\<And>a b. \<lbrakk> ide a; ide b \<rbrakk> \<Longrightarrow> ide (a \<otimes> b)" using tensor_preserves_ide by auto
fix f a b g c d
show "\<lbrakk> \<guillemotleft>f : a \<rightarrow> b\<guillemotright>; \<guillemotleft>g : c \<rightarrow> d\<guillemotright> \<rbrakk> \<Longrightarrow> \<guillemotleft>f \<otimes> g : a \<otimes> c \<rightarrow> b \<otimes> d\<guillemotright>"
using tensor_in_hom by auto
next
text \<open>Naturality of left unitor.\<close>
fix f
assume f: "arr f"
show "\<l>[cod f] \<cdot> (\<I> \<otimes> f) = f \<cdot> \<l>[dom f]"
proof (intro arr_eqI)
show "par (\<l>[cod f] \<cdot> (\<I> \<otimes> f)) (f \<cdot> \<l>[dom f])"
using f by simp
show "\<^bold>\<l>\<^bold>[COD f\<^bold>] \<^bold>\<cdot> (\<^bold>\<I> \<^bold>\<otimes> rep f) \<in> \<l>[cod f] \<cdot> (\<I> \<otimes> f)"
using f by fastforce
show "rep f \<^bold>\<cdot> \<^bold>\<l>\<^bold>[DOM f\<^bold>] \<in> f \<cdot> \<l>[dom f]"
using f by fastforce
show "\<^bold>\<lfloor>\<^bold>\<l>\<^bold>[COD f\<^bold>] \<^bold>\<cdot> (\<^bold>\<I> \<^bold>\<otimes> rep f)\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f \<^bold>\<cdot> \<^bold>\<l>\<^bold>[DOM f\<^bold>]\<^bold>\<rfloor>"
using f by (simp add: Diag_Diagonalize(1) Diagonalize_DOM Diagonalize_COD)
qed
text \<open>Naturality of right unitor.\<close>
show "\<r>[cod f] \<cdot> (f \<otimes> \<I>) = f \<cdot> \<r>[dom f]"
proof (intro arr_eqI)
show "par (\<r>[cod f] \<cdot> (f \<otimes> \<I>)) (f \<cdot> \<r>[dom f])"
using f by simp
show "\<^bold>\<r>\<^bold>[COD f\<^bold>] \<^bold>\<cdot> (rep f \<^bold>\<otimes> \<^bold>\<I>) \<in> \<r>[cod f] \<cdot> (f \<otimes> \<I>)"
using f by fastforce
show "rep f \<^bold>\<cdot> \<^bold>\<r>\<^bold>[DOM f\<^bold>] \<in> f \<cdot> \<r>[dom f]"
using f by fastforce
show "\<^bold>\<lfloor>\<^bold>\<r>\<^bold>[COD f\<^bold>] \<^bold>\<cdot> (rep f \<^bold>\<otimes> \<^bold>\<I>)\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f \<^bold>\<cdot> \<^bold>\<r>\<^bold>[DOM f\<^bold>]\<^bold>\<rfloor>"
using f by (simp add: Diag_Diagonalize(1) Diagonalize_DOM Diagonalize_COD)
qed
next
text \<open>Naturality of associator.\<close>
fix f0 :: "'c arr" and f1 f2
assume f0: "arr f0" and f1: "arr f1" and f2: "arr f2"
show "\<a>[cod f0, cod f1, cod f2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2)
= (f0 \<otimes> f1 \<otimes> f2) \<cdot> \<a>[dom f0, dom f1, dom f2]"
proof (intro arr_eqI)
show 1: "par (\<a>[cod f0, cod f1, cod f2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2))
((f0 \<otimes> f1 \<otimes> f2) \<cdot> \<a>[dom f0, dom f1, dom f2])"
using f0 f1 f2 by force
show "\<^bold>\<a>\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\<cdot> ((rep f0 \<^bold>\<otimes> rep f1) \<^bold>\<otimes> rep f2)
\<in> \<a>[cod f0, cod f1, cod f2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2)"
using f0 f1 f2 by fastforce
show "(rep f0 \<^bold>\<otimes> rep f1 \<^bold>\<otimes> rep f2) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]
\<in> (f0 \<otimes> f1 \<otimes> f2) \<cdot> \<a>[dom f0, dom f1, dom f2]"
using f0 f1 f2 by fastforce
show "\<^bold>\<lfloor>\<^bold>\<a>\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\<cdot> ((rep f0 \<^bold>\<otimes> rep f1) \<^bold>\<otimes> rep f2)\<^bold>\<rfloor>
= \<^bold>\<lfloor>(rep f0 \<^bold>\<otimes> rep f1 \<^bold>\<otimes> rep f2) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>\<^bold>\<a>\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\<cdot> ((rep f0 \<^bold>\<otimes> rep f1) \<^bold>\<otimes> rep f2)\<^bold>\<rfloor>
= \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
proof -
have b0: "\<^bold>\<lfloor>rep (cod f0)\<^bold>\<rfloor> = Cod \<^bold>\<lfloor>rep f0\<^bold>\<rfloor>"
using f0 Cod_Diagonalize_rep by simp
have b1: "\<^bold>\<lfloor>rep (cod f1)\<^bold>\<rfloor> = Cod \<^bold>\<lfloor>rep f1\<^bold>\<rfloor>"
using f1 Cod_Diagonalize_rep by simp
have b2: "\<^bold>\<lfloor>rep (cod f2)\<^bold>\<rfloor> = Cod \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
using f2 Cod_Diagonalize_rep by simp
have "\<^bold>\<lfloor>\<^bold>\<a>\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\<cdot> ((rep f0 \<^bold>\<otimes> rep f1) \<^bold>\<otimes> rep f2)\<^bold>\<rfloor>
= (\<^bold>\<lfloor>rep (cod f0)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (cod f1)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (cod f2)\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor>
(\<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>)"
using f0 f1 f2 using Diag_Diagonalize(1) TensorDiag_assoc by auto
also have "... = \<^bold>\<lfloor>rep (cod f0)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor>
\<^bold>\<lfloor>rep (cod f1)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (cod f2)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
proof -
have "Seq \<^bold>\<lfloor>rep (cod f0)\<^bold>\<rfloor> \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<and> Seq \<^bold>\<lfloor>rep (cod f1)\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<and>
Seq \<^bold>\<lfloor>rep (cod f2)\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
using f0 f1 f2 rep_in_Hom Diagonalize_in_Hom Dom_Diagonalize_rep Cod_Diagonalize_rep
by auto
thus ?thesis
using f0 f1 f2 b0 b1 b2 TensorDiag_in_Hom TensorDiag_preserves_Diag
Diag_Diagonalize Arr_implies_Ide_Dom Arr_implies_Ide_Cod
CompDiag_TensorDiag
by simp
qed
also have "... = \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>rep (cod f0)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f0\<^bold>\<rfloor>"
using f0 b0 CompDiag_Cod_Diag [of "\<^bold>\<lfloor>rep f0\<^bold>\<rfloor>"] Diag_Diagonalize
by simp
moreover have "\<^bold>\<lfloor>rep (cod f1)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f1\<^bold>\<rfloor>"
using f1 b1 CompDiag_Cod_Diag [of "\<^bold>\<lfloor>rep f1\<^bold>\<rfloor>"] Diag_Diagonalize
by simp
moreover have "\<^bold>\<lfloor>rep (cod f2)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
using f2 b2 CompDiag_Cod_Diag [of "\<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"] Diag_Diagonalize
by simp
ultimately show ?thesis by argo
qed
finally show ?thesis by blast
qed
also have "... = \<^bold>\<lfloor>(rep f0 \<^bold>\<otimes> rep f1 \<^bold>\<otimes> rep f2) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]\<^bold>\<rfloor>"
proof -
have a0: "\<^bold>\<lfloor>rep (dom f0)\<^bold>\<rfloor> = Dom \<^bold>\<lfloor>rep f0\<^bold>\<rfloor>"
using f0 Dom_Diagonalize_rep by simp
have a1: "\<^bold>\<lfloor>rep (dom f1)\<^bold>\<rfloor> = Dom \<^bold>\<lfloor>rep f1\<^bold>\<rfloor>"
using f1 Dom_Diagonalize_rep by simp
have a2: "\<^bold>\<lfloor>rep (dom f2)\<^bold>\<rfloor> = Dom \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
using f2 Dom_Diagonalize_rep by simp
have "\<^bold>\<lfloor>(rep f0 \<^bold>\<otimes> rep f1 \<^bold>\<otimes> rep f2) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]\<^bold>\<rfloor>
= (\<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor>
(\<^bold>\<lfloor>rep (dom f0)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f1)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f2)\<^bold>\<rfloor>)"
using f0 f1 f2 using Diag_Diagonalize(1) TensorDiag_assoc by auto
also have "... = \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f0)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f1)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor>
\<^bold>\<lfloor>rep f2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f2)\<^bold>\<rfloor>"
proof -
have "Seq \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f0)\<^bold>\<rfloor> \<and> Seq \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f1)\<^bold>\<rfloor> \<and>
Seq \<^bold>\<lfloor>rep f2\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f2)\<^bold>\<rfloor>"
using f0 f1 f2 rep_in_Hom Diagonalize_in_Hom Dom_Diagonalize_rep Cod_Diagonalize_rep
by auto
thus ?thesis
using f0 f1 f2 a0 a1 a2 TensorDiag_in_Hom TensorDiag_preserves_Diag
Diag_Diagonalize Arr_implies_Ide_Dom Arr_implies_Ide_Cod
CompDiag_TensorDiag
by force
qed
also have "... = \<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>rep f0\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f0)\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f0\<^bold>\<rfloor>"
using f0 a0 CompDiag_Diag_Dom [of "Diagonalize (rep f0)"] Diag_Diagonalize
by simp
moreover have "\<^bold>\<lfloor>rep f1\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f1)\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f1\<^bold>\<rfloor>"
using f1 a1 CompDiag_Diag_Dom [of "Diagonalize (rep f1)"] Diag_Diagonalize
by simp
moreover have "\<^bold>\<lfloor>rep f2\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep (dom f2)\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f2\<^bold>\<rfloor>"
using f2 a2 CompDiag_Diag_Dom [of "Diagonalize (rep f2)"] Diag_Diagonalize
by simp
ultimately show ?thesis by argo
qed
finally show ?thesis by argo
qed
finally show ?thesis by blast
qed
qed
next
text \<open>Tensor preserves composition (interchange).\<close>
fix f g f' g'
show "\<lbrakk> seq g f; seq g' f' \<rbrakk> \<Longrightarrow> (g \<otimes> g') \<cdot> (f \<otimes> f') = g \<cdot> f \<otimes> g' \<cdot> f'"
proof -
assume gf: "seq g f"
assume gf': "seq g' f'"
show ?thesis
proof (intro arr_eqI)
show "par ((g \<otimes> g') \<cdot> (f \<otimes> f')) (g \<cdot> f \<otimes> g' \<cdot> f')"
using gf gf' by fastforce
show "(rep g \<^bold>\<otimes> rep g') \<^bold>\<cdot> (rep f \<^bold>\<otimes> rep f') \<in> (g \<otimes> g') \<cdot> (f \<otimes> f')"
using gf gf' by force
show "rep g \<^bold>\<cdot> rep f \<^bold>\<otimes> rep g' \<^bold>\<cdot> rep f' \<in> g \<cdot> f \<otimes> g' \<cdot> f'"
using gf gf'
by (meson Comp_in_comp_ARR Tensor_in_tensor rep_in_arr seqE seq_char')
show "\<^bold>\<lfloor>(rep g \<^bold>\<otimes> rep g') \<^bold>\<cdot> (rep f \<^bold>\<otimes> rep f')\<^bold>\<rfloor> = \<^bold>\<lfloor>rep g \<^bold>\<cdot> rep f \<^bold>\<otimes> rep g' \<^bold>\<cdot> rep f'\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>(rep g \<^bold>\<otimes> rep g') \<^bold>\<cdot> (rep f \<^bold>\<otimes> rep f')\<^bold>\<rfloor>
= (\<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep g'\<^bold>\<rfloor>) \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> (\<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f'\<^bold>\<rfloor>)"
by auto
also have "... = \<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>rep g'\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f'\<^bold>\<rfloor>"
using gf gf' Arr_rep Diagonalize_rep_preserves_seq
CompDiag_TensorDiag [of "\<^bold>\<lfloor>rep g\<^bold>\<rfloor>" " \<^bold>\<lfloor>rep g'\<^bold>\<rfloor>" "\<^bold>\<lfloor>rep f\<^bold>\<rfloor>" "\<^bold>\<lfloor>rep f'\<^bold>\<rfloor>"]
Diag_Diagonalize Diagonalize_DOM Diagonalize_COD
by force
also have "... = \<^bold>\<lfloor>rep g \<^bold>\<cdot> rep f \<^bold>\<otimes> rep g' \<^bold>\<cdot> rep f'\<^bold>\<rfloor>"
by auto
finally show ?thesis by blast
qed
qed
qed
next
text \<open>The triangle.\<close>
fix a b
assume a: "ide a"
assume b: "ide b"
show "(a \<otimes> \<l>[b]) \<cdot> \<a>[a, \<I>, b] = \<r>[a] \<otimes> b"
proof -
have "par ((a \<otimes> \<l>[b]) \<cdot> \<a>[a, \<I>, b]) (\<r>[a] \<otimes> b)"
using a b by simp
moreover have "can ((a \<otimes> \<l>[b]) \<cdot> \<a>[a, \<I>, b])"
using a b ide_implies_can comp_preserves_can tensor_preserves_can can_assoc can_lunit
by simp
moreover have "can (\<r>[a] \<otimes> b)"
using a b ide_implies_can can_runit tensor_preserves_can by simp
ultimately show ?thesis using can_coherence by blast
qed
next
text \<open>The pentagon.\<close>
fix a b c d
assume a: "ide a"
assume b: "ide b"
assume c: "ide c"
assume d: "ide d"
show "(a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d] \<cdot> (\<a>[a, b, c] \<otimes> d)
= \<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
proof -
let ?LHS = "(a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d] \<cdot> (\<a>[a, b, c] \<otimes> d)"
let ?RHS = "\<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
have "par ?LHS ?RHS"
using a b c d can_assoc tensor_preserves_ide by auto
moreover have "can ?LHS"
using a b c d ide_implies_can comp_preserves_can tensor_preserves_can can_assoc
tensor_preserves_ide
by simp
moreover have "can ?RHS"
using a b c d comp_preserves_can tensor_preserves_can can_assoc tensor_in_hom
tensor_preserves_ide
by simp
ultimately show ?thesis using can_coherence by blast
qed
qed
lemma is_elementary_monoidal_category:
shows "elementary_monoidal_category
comp tensor\<^sub>F\<^sub>M\<^sub>C unity\<^sub>F\<^sub>M\<^sub>C lunit\<^sub>F\<^sub>M\<^sub>C runit\<^sub>F\<^sub>M\<^sub>C assoc\<^sub>F\<^sub>M\<^sub>C"
..
abbreviation T\<^sub>F\<^sub>M\<^sub>C where "T\<^sub>F\<^sub>M\<^sub>C \<equiv> EMC.T"
abbreviation \<alpha>\<^sub>F\<^sub>M\<^sub>C where "\<alpha>\<^sub>F\<^sub>M\<^sub>C \<equiv> EMC.\<alpha>"
abbreviation \<iota>\<^sub>F\<^sub>M\<^sub>C where "\<iota>\<^sub>F\<^sub>M\<^sub>C \<equiv> EMC.\<iota>"
interpretation MC: monoidal_category comp T\<^sub>F\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>M\<^sub>C
using EMC.induces_monoidal_category by auto
lemma induces_monoidal_category:
shows "monoidal_category comp T\<^sub>F\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>M\<^sub>C"
..
end
sublocale free_monoidal_category \<subseteq>
elementary_monoidal_category
comp tensor\<^sub>F\<^sub>M\<^sub>C unity\<^sub>F\<^sub>M\<^sub>C lunit\<^sub>F\<^sub>M\<^sub>C runit\<^sub>F\<^sub>M\<^sub>C assoc\<^sub>F\<^sub>M\<^sub>C
using is_elementary_monoidal_category by auto
sublocale free_monoidal_category \<subseteq> monoidal_category comp T\<^sub>F\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>M\<^sub>C
using induces_monoidal_category by auto
section "Proof of Freeness"
text \<open>
Now we proceed on to establish the freeness of \<open>\<F>C\<close>: each functor
from @{term C} to a monoidal category @{term D} extends uniquely
to a strict monoidal functor from \<open>\<F>C\<close> to D.
\<close>
context free_monoidal_category
begin
lemma rep_lunit:
assumes "ide a"
shows "rep \<l>[a] = \<^bold>\<parallel>\<^bold>\<l>\<^bold>[rep a\<^bold>]\<^bold>\<parallel>"
using assms Lunit_in_lunit [of a "rep a"] rep_in_arr norm_memb_eq_rep [of "\<l>[a]"]
by simp
lemma rep_runit:
assumes "ide a"
shows "rep \<r>[a] = \<^bold>\<parallel>\<^bold>\<r>\<^bold>[rep a\<^bold>]\<^bold>\<parallel>"
using assms Runit_in_runit [of a "rep a"] rep_in_arr norm_memb_eq_rep [of "\<r>[a]"]
by simp
lemma rep_assoc:
assumes "ide a" and "ide b" and "ide c"
shows "rep \<a>[a, b, c] = \<^bold>\<parallel>\<^bold>\<a>\<^bold>[rep a, rep b, rep c\<^bold>]\<^bold>\<parallel>"
using assms Assoc_in_assoc [of a b c "rep a" "rep b" "rep c"] rep_in_arr
norm_memb_eq_rep [of "\<a>[a, b, c]"]
by simp
lemma mkarr_Unity:
shows "mkarr \<^bold>\<I> = \<I>"
using unity\<^sub>F\<^sub>M\<^sub>C_def by simp
text \<open>
The unitors and associator were given syntactic definitions in terms of
corresponding terms, but these were only for the special case of identity
arguments (\emph{i.e.}~the components of the natural transformations).
We need to show that @{term mkarr} gives the correct result for \emph{all}
terms.
\<close>
lemma mkarr_Lunit:
assumes "Arr t"
shows "mkarr \<^bold>\<l>\<^bold>[t\<^bold>] = \<ll> (mkarr t)"
proof -
have "mkarr \<^bold>\<l>\<^bold>[t\<^bold>] = mkarr (t \<^bold>\<cdot> \<^bold>\<l>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>\<^bold>])"
using assms Arr_implies_Ide_Dom Ide_in_Hom Diagonalize_preserves_Ide
Diag_Diagonalize Par_Arr_norm
by (intro mkarr_eqI) simp_all
also have "... = mkarr t \<cdot> mkarr \<^bold>\<l>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>\<^bold>]"
using assms Arr_implies_Ide_Dom Par_Arr_norm Ide_in_Hom by simp
also have "... = mkarr t \<cdot> \<l>[dom (mkarr t)]"
proof -
have "arr \<l>[mkarr (Dom t)]"
using assms Arr_implies_Ide_Dom ide_mkarr_Ide by simp
moreover have "\<^bold>\<l>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>\<^bold>] \<in> \<l>[mkarr (Dom t)]"
using assms Arr_implies_Ide_Dom Lunit_in_lunit rep_mkarr
rep_in_arr [of "mkarr (Dom t)"]
by simp
ultimately show ?thesis
using assms mkarr_memb(2) by simp
qed
also have "... = \<ll> (mkarr t)"
using assms Arr_implies_Ide_Dom ide_mkarr_Ide lunit_agreement by simp
finally show ?thesis by blast
qed
lemma mkarr_Lunit':
assumes "Arr t"
shows "mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<ll>' (mkarr t)"
proof -
have "mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = mkarr (\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>] \<^bold>\<cdot> t)"
using assms Arr_implies_Ide_Cod Ide_in_Hom Diagonalize_preserves_Ide
Diag_Diagonalize Par_Arr_norm
by (intro mkarr_eqI) simp_all
also have "... = mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>] \<cdot> mkarr t"
using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm by simp
also have "... = mkarr (Inv \<^bold>\<l>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>]) \<cdot> mkarr t"
proof -
have "mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>] = mkarr (Inv \<^bold>\<l>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>])"
using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm Inv_in_Hom
Ide_implies_Can norm_preserves_Can Diagonalize_Inv Diagonalize_preserves_Ide
by (intro mkarr_eqI, simp_all)
thus ?thesis by argo
qed
also have "... = \<ll>' (cod (mkarr t)) \<cdot> mkarr t"
proof -
have "mkarr (Inv \<^bold>\<l>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>]) \<cdot> mkarr t = lunit' (cod (mkarr t)) \<cdot> mkarr t"
using assms Arr_implies_Ide_Cod rep_mkarr Par_Arr_norm inv_mkarr
norm_preserves_Can Ide_implies_Can lunit_agreement \<ll>'_ide_simp
Can_implies_Arr arr_mkarr cod_mkarr ide_cod lunit\<^sub>F\<^sub>M\<^sub>C_def
by (metis (no_types, lifting) Can.simps(5))
also have "... = \<ll>' (cod (mkarr t)) \<cdot> mkarr t"
using assms \<ll>'_ide_simp arr_mkarr ide_cod by presburger
finally show ?thesis by blast
qed
also have "... = \<ll>' (mkarr t)"
using assms \<ll>'.is_natural_2 [of "mkarr t"] by simp
finally show ?thesis by blast
qed
lemma mkarr_Runit:
assumes "Arr t"
shows "mkarr \<^bold>\<r>\<^bold>[t\<^bold>] = \<rho> (mkarr t)"
proof -
have "mkarr \<^bold>\<r>\<^bold>[t\<^bold>] = mkarr (t \<^bold>\<cdot> \<^bold>\<r>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>\<^bold>])"
proof -
have "\<not> Diag (Dom t \<^bold>\<otimes> \<^bold>\<I>)" by (cases "Dom t") simp_all
thus ?thesis
using assms Par_Arr_norm Arr_implies_Ide_Dom Ide_in_Hom Diag_Diagonalize
Diagonalize_preserves_Ide
by (intro mkarr_eqI) simp_all
qed
also have "... = mkarr t \<cdot> mkarr \<^bold>\<r>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>\<^bold>]"
using assms Arr_implies_Ide_Dom Par_Arr_norm Ide_in_Hom by simp
also have "... = mkarr t \<cdot> \<r>[dom (mkarr t)]"
proof -
have "arr \<r>[mkarr (Dom t)]"
using assms Arr_implies_Ide_Dom ide_mkarr_Ide by simp
moreover have "\<^bold>\<r>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>\<^bold>] \<in> \<r>[mkarr (Dom t)]"
using assms Arr_implies_Ide_Dom Runit_in_runit rep_mkarr
rep_in_arr [of "mkarr (Dom t)"]
by simp
moreover have "mkarr (Dom t) = mkarr \<^bold>\<parallel>Dom t\<^bold>\<parallel>"
using assms mkarr_rep rep_mkarr arr_mkarr Ide_implies_Arr Arr_implies_Ide_Dom
by metis
ultimately show ?thesis
using assms mkarr_memb(2) by simp
qed
also have "... = \<rho> (mkarr t)"
using assms Arr_implies_Ide_Dom ide_mkarr_Ide runit_agreement by simp
finally show ?thesis by blast
qed
lemma mkarr_Runit':
assumes "Arr t"
shows "mkarr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<rho>' (mkarr t)"
proof -
have "mkarr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = mkarr (\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>] \<^bold>\<cdot> t)"
proof -
have "\<not> Diag (Cod t \<^bold>\<otimes> \<^bold>\<I>)" by (cases "Cod t") simp_all
thus ?thesis
using assms Par_Arr_norm Arr_implies_Ide_Cod Ide_in_Hom
Diagonalize_preserves_Ide Diag_Diagonalize
by (intro mkarr_eqI) simp_all
qed
also have "... = mkarr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>] \<cdot> mkarr t"
using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm by simp
also have "... = mkarr (Inv \<^bold>\<r>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>]) \<cdot> mkarr t"
proof -
have "mkarr (Runit' (norm (Cod t))) = mkarr (Inv (Runit (norm (Cod t))))"
using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm Inv_in_Hom
Ide_implies_Can norm_preserves_Can Diagonalize_Inv Diagonalize_preserves_Ide
by (intro mkarr_eqI) simp_all
thus ?thesis by argo
qed
also have "... = \<rho>' (cod (mkarr t)) \<cdot> mkarr t"
proof -
have "mkarr (Inv \<^bold>\<r>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>\<^bold>]) \<cdot> mkarr t = runit' (cod (mkarr t)) \<cdot> mkarr t"
using assms Arr_implies_Ide_Cod rep_mkarr inv_mkarr norm_preserves_Can
Ide_implies_Can runit_agreement Can_implies_Arr arr_mkarr cod_mkarr
ide_cod runit\<^sub>F\<^sub>M\<^sub>C_def
by (metis (no_types, lifting) Can.simps(7))
also have "... = \<rho>' (cod (mkarr t)) \<cdot> mkarr t"
proof -
have "runit' (cod (mkarr t)) = \<rho>' (cod (mkarr t))"
using assms \<rho>'_ide_simp arr_mkarr ide_cod by blast
thus ?thesis by argo
qed
finally show ?thesis by blast
qed
also have "... = \<rho>' (mkarr t)"
using assms \<rho>'.is_natural_2 [of "mkarr t"] by simp
finally show ?thesis by blast
qed
lemma mkarr_Assoc:
assumes "Arr t" and "Arr u" and "Arr v"
shows "mkarr \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = \<alpha> (mkarr t, mkarr u, mkarr v)"
proof -
have "mkarr \<^bold>\<a>\<^bold>[t, u, v\<^bold>] = mkarr ((t \<^bold>\<otimes> u \<^bold>\<otimes> v) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<parallel>Dom t\<^bold>\<parallel>, \<^bold>\<parallel>Dom u\<^bold>\<parallel>, \<^bold>\<parallel>Dom v\<^bold>\<parallel>\<^bold>])"
using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Ide_in_Hom
Diag_Diagonalize Diagonalize_preserves_Ide TensorDiag_preserves_Ide
TensorDiag_preserves_Diag TensorDiag_assoc Par_Arr_norm
by (intro mkarr_eqI, simp_all)
also have "... = \<alpha> (mkarr t, mkarr u, mkarr v)"
using assms Arr_implies_Ide_Dom rep_mkarr Ide_in_Hom assoc\<^sub>F\<^sub>M\<^sub>C_def
Par_Arr_norm [of "Dom t"] Par_Arr_norm [of "Dom u"] Par_Arr_norm [of "Dom v"]
\<alpha>_simp
by simp
finally show ?thesis by blast
qed
lemma mkarr_Assoc':
assumes "Arr t" and "Arr u" and "Arr v"
shows "mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = \<alpha>' (mkarr t, mkarr u, mkarr v)"
proof -
have "mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = mkarr (\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>, \<^bold>\<parallel>Cod u\<^bold>\<parallel>, \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>] \<^bold>\<cdot> (t \<^bold>\<otimes> u \<^bold>\<otimes> v))"
using assms Par_Arr_norm Arr_implies_Ide_Cod Ide_in_Hom Diag_Diagonalize
TensorDiag_preserves_Diag CompDiag_Cod_Diag [of "\<^bold>\<lfloor>t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>v\<^bold>\<rfloor>"]
by (intro mkarr_eqI, simp_all)
also have "... = mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>, \<^bold>\<parallel>Cod u\<^bold>\<parallel>, \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>] \<cdot> mkarr (t \<^bold>\<otimes> u \<^bold>\<otimes> v)"
using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm by simp
also have "... = mkarr (Inv \<^bold>\<a>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>, \<^bold>\<parallel>Cod u\<^bold>\<parallel>, \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>]) \<cdot> mkarr (t \<^bold>\<otimes> u \<^bold>\<otimes> v)"
proof -
have "mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>, \<^bold>\<parallel>Cod u\<^bold>\<parallel>, \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>] =
mkarr (Inv \<^bold>\<a>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>, \<^bold>\<parallel>Cod u\<^bold>\<parallel>, \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>])"
using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm Inv_in_Hom Ide_implies_Can
norm_preserves_Can Diagonalize_Inv Diagonalize_preserves_Ide
by (intro mkarr_eqI, simp_all)
thus ?thesis by argo
qed
also have "... = inv (mkarr \<^bold>\<a>\<^bold>[\<^bold>\<parallel>Cod t\<^bold>\<parallel>, \<^bold>\<parallel>Cod u\<^bold>\<parallel>, \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>]) \<cdot> mkarr (t \<^bold>\<otimes> u \<^bold>\<otimes> v)"
using assms Arr_implies_Ide_Cod Ide_implies_Can norm_preserves_Can by simp
also have "... = \<alpha>' (mkarr t, mkarr u, mkarr v)"
proof -
have "mkarr (\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Inv \<^bold>\<parallel>Cod t\<^bold>\<parallel>, Inv \<^bold>\<parallel>Cod u\<^bold>\<parallel>, Inv \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>] \<^bold>\<cdot> (Cod t \<^bold>\<otimes> Cod u \<^bold>\<otimes> Cod v))
= mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[Inv \<^bold>\<parallel>Cod t\<^bold>\<parallel>, Inv \<^bold>\<parallel>Cod u\<^bold>\<parallel>, Inv \<^bold>\<parallel>Cod v\<^bold>\<parallel>\<^bold>]"
using assms Arr_implies_Ide_Cod Inv_in_Hom norm_preserves_Can Diagonalize_Inv
Ide_implies_Can Diag_Diagonalize Ide_in_Hom Diagonalize_preserves_Ide
Par_Arr_norm TensorDiag_preserves_Diag
CompDiag_Cod_Diag [of "\<^bold>\<lfloor>Cod t\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod u\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>Cod v\<^bold>\<rfloor>"]
by (intro mkarr_eqI) simp_all
thus ?thesis
using assms Arr_implies_Ide_Cod rep_mkarr assoc\<^sub>F\<^sub>M\<^sub>C_def \<alpha>'.map_simp by simp
qed
finally show ?thesis by blast
qed
text \<open>
Next, we define the ``inclusion of generators'' functor from @{term C} to \<open>\<F>C\<close>.
\<close>
definition inclusion_of_generators
where "inclusion_of_generators \<equiv> \<lambda>f. if C.arr f then mkarr \<^bold>\<langle>f\<^bold>\<rangle> else null"
lemma inclusion_is_functor:
shows "functor C comp inclusion_of_generators"
unfolding inclusion_of_generators_def
apply unfold_locales
apply auto[4]
by (elim C.seqE, simp, intro mkarr_eqI, auto)
end
text \<open>
We now show that, given a functor @{term V} from @{term C} to a
a monoidal category @{term D}, the evaluation map that takes formal arrows
of the monoidal language of @{term C} to arrows of @{term D}
induces a strict monoidal functor from \<open>\<F>C\<close> to @{term D}.
\<close>
locale evaluation_functor =
C: category C +
D: monoidal_category D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D +
evaluation_map C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V +
\<F>C: free_monoidal_category C
for C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and T\<^sub>D :: "'d * 'd \<Rightarrow> 'd"
and \<alpha>\<^sub>D :: "'d * 'd * 'd \<Rightarrow> 'd"
and \<iota>\<^sub>D :: "'d"
and V :: "'c \<Rightarrow> 'd"
begin
notation eval ("\<lbrace>_\<rbrace>")
definition map
where "map f \<equiv> if \<F>C.arr f then \<lbrace>\<F>C.rep f\<rbrace> else D.null"
text \<open>
It follows from the coherence theorem that a formal arrow and its normal
form always have the same evaluation.
\<close>
lemma eval_norm:
assumes "Arr t"
shows "\<lbrace>\<^bold>\<parallel>t\<^bold>\<parallel>\<rbrace> = \<lbrace>t\<rbrace>"
using assms \<F>C.Par_Arr_norm \<F>C.Diagonalize_norm coherence canonical_factorization
by simp
interpretation "functor" \<F>C.comp D map
proof
fix f
show "\<not>\<F>C.arr f \<Longrightarrow> map f = D.null" using map_def by simp
assume f: "\<F>C.arr f"
show "D.arr (map f)" using f map_def \<F>C.arr_char by simp
show "D.dom (map f) = map (\<F>C.dom f)"
using f map_def eval_norm \<F>C.rep_dom Arr_implies_Ide_Dom by auto
show "D.cod (map f) = map (\<F>C.cod f)"
using f map_def eval_norm \<F>C.rep_cod Arr_implies_Ide_Cod by auto
next
fix f g
assume fg: "\<F>C.seq g f"
show "map (\<F>C.comp g f) = D (map g) (map f)"
using fg map_def \<F>C.rep_comp \<F>C.rep_preserves_seq eval_norm by auto
qed
lemma is_functor:
shows "functor \<F>C.comp D map" ..
interpretation FF: product_functor \<F>C.comp \<F>C.comp D D map map ..
interpretation FoT: composite_functor \<F>C.CC.comp \<F>C.comp D \<F>C.T\<^sub>F\<^sub>M\<^sub>C map ..
interpretation ToFF: composite_functor \<F>C.CC.comp D.CC.comp D FF.map T\<^sub>D ..
interpretation strict_monoidal_functor
\<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D map
proof
show "map \<F>C.\<iota> = \<iota>\<^sub>D"
using \<F>C.\<iota>_def \<F>C.lunit_agreement map_def \<F>C.rep_lunit \<F>C.Arr_rep [of \<I>]
eval_norm \<F>C.lunit_agreement D.unitor_coincidence D.comp_cod_arr D.\<iota>_in_hom
by auto
show "\<And>f g. \<lbrakk> \<F>C.arr f; \<F>C.arr g \<rbrakk> \<Longrightarrow>
map (\<F>C.tensor f g) = D.tensor (map f) (map g)"
using map_def \<F>C.rep_tensor \<F>C.Arr_rep eval_norm by simp
show "\<And>a b c. \<lbrakk> \<F>C.ide a; \<F>C.ide b; \<F>C.ide c \<rbrakk> \<Longrightarrow>
map (\<F>C.assoc a b c) = D.assoc (map a) (map b) (map c)"
using map_def \<F>C.assoc\<^sub>F\<^sub>M\<^sub>C_def \<F>C.rep_mkarr eval_norm by auto
qed
lemma is_strict_monoidal_functor:
shows "strict_monoidal_functor \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D map"
..
end
sublocale evaluation_functor \<subseteq> strict_monoidal_functor
\<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha>\<^sub>F\<^sub>M\<^sub>C \<F>C.\<iota>\<^sub>F\<^sub>M\<^sub>C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D map
using is_strict_monoidal_functor by auto
text \<open>
The final step in proving freeness is to show that the evaluation functor
is the \emph{unique} strict monoidal extension of the functor @{term V}
to \<open>\<F>C\<close>. This is done by induction, exploiting the syntactic construction
of \<open>\<F>C\<close>.
\<close>
text \<open>
To ease the statement and proof of the result, we define a locale that
expresses that @{term F} is a strict monoidal extension to monoidal
category @{term C}, of a functor @{term "V"} from @{term "C\<^sub>0"} to a
monoidal category @{term D}, along a functor @{term I} from
@{term "C\<^sub>0"} to @{term C}.
\<close>
locale strict_monoidal_extension =
C\<^sub>0: category C\<^sub>0 +
C: monoidal_category C T\<^sub>C \<alpha>\<^sub>C \<iota>\<^sub>C +
D: monoidal_category D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D +
I: "functor" C\<^sub>0 C I +
V: "functor" C\<^sub>0 D V +
strict_monoidal_functor C T\<^sub>C \<alpha>\<^sub>C \<iota>\<^sub>C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D F
for C\<^sub>0 :: "'c\<^sub>0 comp"
and C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and T\<^sub>C :: "'c * 'c \<Rightarrow> 'c"
and \<alpha>\<^sub>C :: "'c * 'c * 'c \<Rightarrow> 'c"
and \<iota>\<^sub>C :: "'c"
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and T\<^sub>D :: "'d * 'd \<Rightarrow> 'd"
and \<alpha>\<^sub>D :: "'d * 'd * 'd \<Rightarrow> 'd"
and \<iota>\<^sub>D :: "'d"
and I :: "'c\<^sub>0 \<Rightarrow> 'c"
and V :: "'c\<^sub>0 \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd" +
assumes is_extension: "\<forall>f. C\<^sub>0.arr f \<longrightarrow> F (I f) = V f"
sublocale evaluation_functor \<subseteq>
strict_monoidal_extension C \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
\<F>C.inclusion_of_generators V map
proof -
interpret inclusion: "functor" C \<F>C.comp \<F>C.inclusion_of_generators
using \<F>C.inclusion_is_functor by auto
show "strict_monoidal_extension C \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
\<F>C.inclusion_of_generators V map"
apply unfold_locales
using map_def \<F>C.rep_mkarr eval_norm \<F>C.inclusion_of_generators_def by simp
qed
text \<open>
A special case of interest is a strict monoidal extension to \<open>\<F>C\<close>,
of a functor @{term V} from a category @{term C} to a monoidal category @{term D},
along the inclusion of generators from @{term C} to \<open>\<F>C\<close>.
The evaluation functor induced by @{term V} is such an extension.
\<close>
locale strict_monoidal_extension_to_free_monoidal_category =
C: category C +
monoidal_language C +
\<F>C: free_monoidal_category C +
strict_monoidal_extension C \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
\<F>C.inclusion_of_generators V F
for C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and T\<^sub>D :: "'d * 'd \<Rightarrow> 'd"
and \<alpha>\<^sub>D :: "'d * 'd * 'd \<Rightarrow> 'd"
and \<iota>\<^sub>D :: "'d"
and V :: "'c \<Rightarrow> 'd"
and F :: "'c free_monoidal_category.arr \<Rightarrow> 'd"
begin
lemma strictly_preserves_everything:
shows "C.arr f \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<langle>f\<^bold>\<rangle>) = V f"
and "F (\<F>C.mkarr \<^bold>\<I>) = \<I>\<^sub>D"
and "\<lbrakk> Arr t; Arr u \<rbrakk> \<Longrightarrow> F (\<F>C.mkarr (t \<^bold>\<otimes> u)) = F (\<F>C.mkarr t) \<otimes>\<^sub>D F (\<F>C.mkarr u)"
and "\<lbrakk> Arr t; Arr u; Dom t = Cod u \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr (t \<^bold>\<cdot> u)) = F (\<F>C.mkarr t) \<cdot>\<^sub>D F (\<F>C.mkarr u)"
and "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<l>\<^bold>[t\<^bold>]) = D.\<ll> (F (\<F>C.mkarr t))"
and "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\<ll>'.map (F (\<F>C.mkarr t))"
and "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<r>\<^bold>[t\<^bold>]) = D.\<rho> (F (\<F>C.mkarr t))"
and "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\<rho>'.map (F (\<F>C.mkarr t))"
and "\<lbrakk> Arr t; Arr u; Arr v \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr \<^bold>\<a>\<^bold>[t, u, v\<^bold>]) = \<alpha>\<^sub>D (F (\<F>C.mkarr t), F (\<F>C.mkarr u), F (\<F>C.mkarr v))"
and "\<lbrakk> Arr t; Arr u; Arr v \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>])
= D.\<alpha>' (F (\<F>C.mkarr t), F (\<F>C.mkarr u), F (\<F>C.mkarr v))"
proof -
show "C.arr f \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<langle>f\<^bold>\<rangle>) = V f"
using is_extension \<F>C.inclusion_of_generators_def by simp
show "F (\<F>C.mkarr \<^bold>\<I>) = \<I>\<^sub>D"
using \<F>C.mkarr_Unity \<F>C.\<iota>_def strictly_preserves_unity \<F>C.\<I>_agreement by auto
show tensor_case:
"\<And>t u.\<lbrakk> Arr t; Arr u \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr (t \<^bold>\<otimes> u)) = F (\<F>C.mkarr t) \<otimes>\<^sub>D F (\<F>C.mkarr u)"
proof -
fix t u
assume t: "Arr t" and u: "Arr u"
have "F (\<F>C.mkarr (t \<^bold>\<otimes> u)) = F (\<F>C.tensor (\<F>C.mkarr t) (\<F>C.mkarr u))"
using t u \<F>C.tensor_mkarr \<F>C.arr_mkarr by simp
also have "... = F (\<F>C.mkarr t) \<otimes>\<^sub>D F (\<F>C.mkarr u)"
using t u \<F>C.arr_mkarr strictly_preserves_tensor by blast
finally show "F (\<F>C.mkarr (t \<^bold>\<otimes> u)) = F (\<F>C.mkarr t) \<otimes>\<^sub>D F (\<F>C.mkarr u)"
by fast
qed
show "\<lbrakk> Arr t; Arr u; Dom t = Cod u \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr (t \<^bold>\<cdot> u)) = F (\<F>C.mkarr t) \<cdot>\<^sub>D F (\<F>C.mkarr u)"
proof -
fix t u
assume t: "Arr t" and u: "Arr u" and tu: "Dom t = Cod u"
show "F (\<F>C.mkarr (t \<^bold>\<cdot> u)) = F (\<F>C.mkarr t) \<cdot>\<^sub>D F (\<F>C.mkarr u)"
proof -
have "F (\<F>C.mkarr (t \<^bold>\<cdot> u)) = F (\<F>C.mkarr t \<cdot> \<F>C.mkarr u)"
using t u tu \<F>C.comp_mkarr by simp
also have "... = F (\<F>C.mkarr t) \<cdot>\<^sub>D F (\<F>C.mkarr u)"
using t u tu \<F>C.arr_mkarr by fastforce
finally show ?thesis by blast
qed
qed
show "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<l>\<^bold>[t\<^bold>]) = D.\<ll> (F (\<F>C.mkarr t))"
using \<F>C.mkarr_Lunit Arr_implies_Ide_Dom \<F>C.ide_mkarr_Ide strictly_preserves_lunit
by simp
show "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<r>\<^bold>[t\<^bold>]) = D.\<rho> (F (\<F>C.mkarr t))"
using \<F>C.mkarr_Runit Arr_implies_Ide_Dom \<F>C.ide_mkarr_Ide strictly_preserves_runit
by simp
show "\<lbrakk> Arr t; Arr u; Arr v \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr \<^bold>\<a>\<^bold>[t, u, v\<^bold>])
= \<alpha>\<^sub>D (F (\<F>C.mkarr t), F (\<F>C.mkarr u), F (\<F>C.mkarr v))"
using \<F>C.mkarr_Assoc strictly_preserves_assoc \<F>C.ide_mkarr_Ide tensor_case
by simp
show "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\<ll>'.map (F (\<F>C.mkarr t))"
proof -
assume t: "Arr t"
have "F (\<F>C.mkarr \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = F (\<F>C.lunit' (\<F>C.mkarr (Cod t))) \<cdot>\<^sub>D F (\<F>C.mkarr t)"
using t \<F>C.mkarr_Lunit' Arr_implies_Ide_Cod \<F>C.ide_mkarr_Ide \<F>C.\<ll>'.map_simp
\<F>C.comp_cod_arr
by simp
also have "... = D.lunit' (D.cod (F (\<F>C.mkarr t))) \<cdot>\<^sub>D F (\<F>C.mkarr t)"
using t Arr_implies_Ide_Cod \<F>C.ide_mkarr_Ide strictly_preserves_lunit
preserves_inv
by simp
also have "... = D.\<ll>'.map (F (\<F>C.mkarr t))"
using t D.\<ll>'.map_simp D.comp_cod_arr by simp
finally show ?thesis by blast
qed
show "Arr t \<Longrightarrow> F (\<F>C.mkarr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\<rho>'.map (F (\<F>C.mkarr t))"
proof -
assume t: "Arr t"
have "F (\<F>C.mkarr \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = F (\<F>C.runit' (\<F>C.mkarr (Cod t))) \<cdot>\<^sub>D F (\<F>C.mkarr t)"
using t \<F>C.mkarr_Runit' Arr_implies_Ide_Cod \<F>C.ide_mkarr_Ide \<F>C.\<rho>'.map_simp
\<F>C.comp_cod_arr
by simp
also have "... = D.runit' (D.cod (F (\<F>C.mkarr t))) \<cdot>\<^sub>D F (\<F>C.mkarr t)"
using t Arr_implies_Ide_Cod \<F>C.ide_mkarr_Ide strictly_preserves_runit
preserves_inv
by simp
also have "... = D.\<rho>'.map (F (\<F>C.mkarr t))"
using t D.\<rho>'.map_simp D.comp_cod_arr by simp
finally show ?thesis by blast
qed
show "\<lbrakk> Arr t; Arr u; Arr v \<rbrakk> \<Longrightarrow>
F (\<F>C.mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>])
= D.\<alpha>'.map (F (\<F>C.mkarr t), F (\<F>C.mkarr u), F (\<F>C.mkarr v))"
proof -
assume t: "Arr t" and u: "Arr u" and v: "Arr v"
have "F (\<F>C.mkarr \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]) =
F (\<F>C.assoc' (\<F>C.mkarr (Cod t)) (\<F>C.mkarr (Cod u)) (\<F>C.mkarr (Cod v))) \<cdot>\<^sub>D
(F (\<F>C.mkarr t) \<otimes>\<^sub>D F (\<F>C.mkarr u) \<otimes>\<^sub>D F (\<F>C.mkarr v))"
using t u v \<F>C.mkarr_Assoc' Arr_implies_Ide_Cod \<F>C.ide_mkarr_Ide \<F>C.\<alpha>'.map_simp
tensor_case \<F>C.iso_assoc
by simp
also have "... = D.assoc' (D.cod (F (\<F>C.mkarr t))) (D.cod (F (\<F>C.mkarr u)))
(D.cod (F (\<F>C.mkarr v))) \<cdot>\<^sub>D
(F (\<F>C.mkarr t) \<otimes>\<^sub>D F (\<F>C.mkarr u) \<otimes>\<^sub>D F (\<F>C.mkarr v))"
using t u v \<F>C.ide_mkarr_Ide Arr_implies_Ide_Cod preserves_inv \<F>C.iso_assoc
strictly_preserves_assoc
[of "\<F>C.mkarr (Cod t)" "\<F>C.mkarr (Cod u)" "\<F>C.mkarr (Cod v)"]
by simp
also have "... = D.\<alpha>'.map (F (\<F>C.mkarr t), F (\<F>C.mkarr u), F (\<F>C.mkarr v))"
using t u v D.\<alpha>'.map_simp by simp
finally show ?thesis by blast
qed
qed
end
sublocale evaluation_functor \<subseteq> strict_monoidal_extension_to_free_monoidal_category
C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V map
..
context free_monoidal_category
begin
text \<open>
The evaluation functor induced by @{term V} is the unique strict monoidal
extension of @{term V} to \<open>\<F>C\<close>.
\<close>
theorem is_free:
assumes "strict_monoidal_extension_to_free_monoidal_category C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V F"
shows "F = evaluation_functor.map C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V"
proof -
interpret F: strict_monoidal_extension_to_free_monoidal_category C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V F
using assms by auto
interpret E: evaluation_functor C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V ..
have Ide_case: "\<And>a. Ide a \<Longrightarrow> F (mkarr a) = E.map (mkarr a)"
proof -
fix a
show "Ide a \<Longrightarrow> F (mkarr a) = E.map (mkarr a)"
using E.strictly_preserves_everything F.strictly_preserves_everything Ide_implies_Arr
by (induct a) auto
qed
show ?thesis
proof
fix f
have "\<not>arr f \<Longrightarrow> F f = E.map f"
using E.is_extensional F.is_extensional by simp
moreover have "arr f \<Longrightarrow> F f = E.map f"
proof -
assume f: "arr f"
have "Arr (rep f) \<and> f = mkarr (rep f)" using f mkarr_rep by simp
moreover have "\<And>t. Arr t \<Longrightarrow> F (mkarr t) = E.map (mkarr t)"
proof -
fix t
show "Arr t \<Longrightarrow> F (mkarr t) = E.map (mkarr t)"
using Ide_case E.strictly_preserves_everything F.strictly_preserves_everything
Arr_implies_Ide_Dom Arr_implies_Ide_Cod
by (induct t) auto
qed
ultimately show "F f = E.map f" by metis
qed
ultimately show "F f = E.map f" by blast
qed
qed
end
section "Strict Subcategory"
context free_monoidal_category
begin
text \<open>
In this section we show that \<open>\<F>C\<close> is monoidally equivalent to its full subcategory
\<open>\<F>\<^sub>SC\<close> whose objects are the equivalence classes of diagonal identity terms,
and that this subcategory is the free strict monoidal category generated by @{term C}.
\<close>
interpretation \<F>\<^sub>SC: full_subcategory comp \<open>\<lambda>f. ide f \<and> Diag (DOM f)\<close>
by (unfold_locales) auto
text \<open>
The mapping defined on equivalence classes by diagonalizing their representatives
is a functor from the free monoidal category to the subcategory @{term "\<F>\<^sub>SC"}.
\<close>
definition D
where "D \<equiv> \<lambda>f. if arr f then mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor> else \<F>\<^sub>SC.null"
text \<open>
The arrows of \<open>\<F>\<^sub>SC\<close> are those equivalence classes whose canonical representative
term has diagonal formal domain and codomain.
\<close>
lemma strict_arr_char:
shows "\<F>\<^sub>SC.arr f \<longleftrightarrow> arr f \<and> Diag (DOM f) \<and> Diag (COD f)"
proof
show "arr f \<and> Diag (DOM f) \<and> Diag (COD f) \<Longrightarrow> \<F>\<^sub>SC.arr f"
- using \<F>\<^sub>SC.arr_char DOM_dom DOM_cod by simp
+ using \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C DOM_dom DOM_cod by simp
show "\<F>\<^sub>SC.arr f \<Longrightarrow> arr f \<and> Diag (DOM f) \<and> Diag (COD f)"
- using \<F>\<^sub>SC.arr_char Arr_rep Arr_implies_Ide_Cod Ide_implies_Arr DOM_dom DOM_cod
+ using \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C Arr_rep Arr_implies_Ide_Cod Ide_implies_Arr DOM_dom DOM_cod
by force
qed
text \<open>
Alternatively, the arrows of \<open>\<F>\<^sub>SC\<close> are those equivalence classes
that are preserved by diagonalization of representatives.
\<close>
lemma strict_arr_char':
shows "\<F>\<^sub>SC.arr f \<longleftrightarrow> arr f \<and> D f = f"
proof
fix f
assume f: "\<F>\<^sub>SC.arr f"
show "arr f \<and> D f = f"
proof
- show "arr f" using f \<F>\<^sub>SC.arr_char by blast
+ show "arr f" using f \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C by blast
show "D f = f"
using f strict_arr_char mkarr_Diagonalize_rep D_def by simp
qed
next
assume f: "arr f \<and> D f = f"
show "\<F>\<^sub>SC.arr f"
proof -
have "arr f" using f by simp
moreover have "Diag (DOM f)"
proof -
have "DOM f = DOM (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)" using f D_def by auto
also have "... = Dom \<^bold>\<parallel>\<^bold>\<lfloor>rep f\<^bold>\<rfloor>\<^bold>\<parallel>"
using f Arr_rep Diagonalize_in_Hom rep_mkarr by simp
also have "... = Dom \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using f Arr_rep Diagonalize_in_Hom Par_Arr_norm [of "\<^bold>\<lfloor>rep f\<^bold>\<rfloor>"] by force
finally have "DOM f = Dom \<^bold>\<lfloor>rep f\<^bold>\<rfloor>" by blast
thus ?thesis using f Arr_rep Diag_Diagonalize Dom_preserves_Diag by metis
qed
moreover have "Diag (COD f)"
proof -
have "COD f = COD (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)" using f D_def by auto
also have "... = Cod \<^bold>\<parallel>\<^bold>\<lfloor>rep f\<^bold>\<rfloor>\<^bold>\<parallel>"
using f Arr_rep Diagonalize_in_Hom rep_mkarr by simp
also have "... = Cod \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using f Arr_rep Diagonalize_in_Hom Par_Arr_norm [of "\<^bold>\<lfloor>rep f\<^bold>\<rfloor>"] by force
finally have "COD f = Cod \<^bold>\<lfloor>rep f\<^bold>\<rfloor>" by blast
thus ?thesis using f Arr_rep Diag_Diagonalize Cod_preserves_Diag by metis
qed
ultimately show ?thesis using strict_arr_char by auto
qed
qed
interpretation D: "functor" comp \<F>\<^sub>SC.comp D
proof -
have 1: "\<And>f. arr f \<Longrightarrow> \<F>\<^sub>SC.arr (D f)"
unfolding strict_arr_char D_def
using arr_mkarr Diagonalize_in_Hom Arr_rep rep_mkarr Par_Arr_norm
Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diag_Diagonalize
by force
show "functor comp \<F>\<^sub>SC.comp D"
proof
show "\<And>f. \<not> arr f \<Longrightarrow> D f = \<F>\<^sub>SC.null" using D_def by simp
show "\<And>f. arr f \<Longrightarrow> \<F>\<^sub>SC.arr (D f)" by fact
show "\<And>f. arr f \<Longrightarrow> \<F>\<^sub>SC.dom (D f) = D (dom f)"
- using D_def Diagonalize_in_Hom \<F>\<^sub>SC.dom_char \<F>\<^sub>SC.arr_char
+ using D_def Diagonalize_in_Hom \<F>\<^sub>SC.dom_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C
rep_mkarr rep_dom Arr_implies_Ide_Dom Arr_implies_Ide_Cod
Diagonalize_preserves_Ide ide_mkarr_Ide Diag_Diagonalize Dom_norm
by simp
show 2: "\<And>f. arr f \<Longrightarrow> \<F>\<^sub>SC.cod (D f) = D (cod f)"
- using D_def Diagonalize_in_Hom \<F>\<^sub>SC.cod_char \<F>\<^sub>SC.arr_char
+ using D_def Diagonalize_in_Hom \<F>\<^sub>SC.cod_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C
rep_mkarr rep_cod Arr_implies_Ide_Dom Arr_implies_Ide_Cod
Diagonalize_preserves_Ide ide_mkarr_Ide Diag_Diagonalize Dom_norm
by simp
fix f g
assume fg: "seq g f"
hence fg': "arr f \<and> arr g \<and> dom g = cod f" by blast
show "D (g \<cdot> f) = \<F>\<^sub>SC.comp (D g) (D f)"
proof -
have seq: "\<F>\<^sub>SC.seq (mkarr \<^bold>\<lfloor>rep g\<^bold>\<rfloor>) (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)"
proof -
have 3: "\<F>\<^sub>SC.arr (mkarr \<^bold>\<lfloor>rep g\<^bold>\<rfloor>) \<and> \<F>\<^sub>SC.arr (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)"
using fg' 1 arr_char D_def by force
moreover have "\<F>\<^sub>SC.dom (mkarr \<^bold>\<lfloor>rep g\<^bold>\<rfloor>) = \<F>\<^sub>SC.cod (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)"
- using fg' 2 3 \<F>\<^sub>SC.dom_char rep_in_Hom mkarr_in_hom D_def
- Dom_Diagonalize_rep Diag_implies_Arr Diag_Diagonalize(1) \<F>\<^sub>SC.arr_char
+ using fg' 2 3 \<F>\<^sub>SC.dom_char\<^sub>S\<^sub>b\<^sub>C rep_in_Hom mkarr_in_hom D_def
+ Dom_Diagonalize_rep Diag_implies_Arr Diag_Diagonalize(1) \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C
by force
ultimately show ?thesis using \<F>\<^sub>SC.seqI by auto
qed
have "mkarr \<^bold>\<lfloor>rep (g \<cdot> f)\<^bold>\<rfloor> = \<F>\<^sub>SC.comp (mkarr \<^bold>\<lfloor>rep g\<^bold>\<rfloor>) (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)"
proof -
have Seq: "Seq \<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using fg rep_preserves_seq Diagonalize_in_Hom by force
hence 4: "\<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<cdot> \<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<in> Hom \<^bold>\<lfloor>DOM f\<^bold>\<rfloor> \<^bold>\<lfloor>COD g\<^bold>\<rfloor>"
using fg' Seq Diagonalize_in_Hom by auto
have "\<F>\<^sub>SC.comp (mkarr \<^bold>\<lfloor>rep g\<^bold>\<rfloor>) (mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>) = mkarr \<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<cdot> mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
- using seq \<F>\<^sub>SC.comp_char \<F>\<^sub>SC.seq_char by meson
+ using seq \<F>\<^sub>SC.comp_char \<F>\<^sub>SC.seq_char\<^sub>S\<^sub>b\<^sub>C by meson
also have "... = mkarr (\<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<cdot> \<^bold>\<lfloor>rep f\<^bold>\<rfloor>)"
using Seq comp_mkarr by fastforce
also have "... = mkarr \<^bold>\<lfloor>rep (g \<cdot> f)\<^bold>\<rfloor>"
proof (intro mkarr_eqI)
show "Par (\<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<cdot> \<^bold>\<lfloor>rep f\<^bold>\<rfloor>) \<^bold>\<lfloor>rep (g \<cdot> f)\<^bold>\<rfloor>"
using fg 4 rep_in_Hom rep_preserves_seq rep_in_Hom Diagonalize_in_Hom
Par_Arr_norm
apply (elim seqE, auto)
by (simp_all add: rep_comp)
show "\<^bold>\<lfloor>\<^bold>\<lfloor>rep g\<^bold>\<rfloor> \<^bold>\<cdot> \<^bold>\<lfloor>rep f\<^bold>\<rfloor>\<^bold>\<rfloor> = \<^bold>\<lfloor>\<^bold>\<lfloor>rep (g \<cdot> f)\<^bold>\<rfloor>\<^bold>\<rfloor>"
using fg rep_preserves_seq norm_in_Hom Diag_Diagonalize Diagonalize_Diag
apply auto
by (simp add: rep_comp)
qed
finally show ?thesis by blast
qed
thus ?thesis using fg D_def by auto
qed
qed
qed
lemma diagonalize_is_functor:
shows "functor comp \<F>\<^sub>SC.comp D" ..
lemma diagonalize_strict_arr:
assumes "\<F>\<^sub>SC.arr f"
shows "D f = f"
using assms arr_char D_def strict_arr_char Arr_rep Arr_implies_Ide_Dom Ide_implies_Arr
mkarr_Diagonalize_rep [of f]
by auto
lemma diagonalize_is_idempotent:
shows "D o D = D"
unfolding D_def
using D.is_extensional \<F>\<^sub>SC.null_char Arr_rep Diagonalize_in_Hom mkarr_Diagonalize_rep
strict_arr_char rep_mkarr
by fastforce
lemma diagonalize_tensor:
assumes "arr f" and "arr g"
shows "D (f \<otimes> g) = D (D f \<otimes> D g)"
unfolding D_def
using assms strict_arr_char rep_in_Hom Diagonalize_in_Hom tensor_mkarr rep_tensor
Diagonalize_in_Hom rep_mkarr Diagonalize_norm Diagonalize_Tensor
by force
lemma ide_diagonalize_can:
assumes "can f"
shows "ide (D f)"
using assms D_def Can_rep_can Ide_Diagonalize_Can ide_mkarr_Ide can_implies_arr
by simp
text \<open>
We next show that the diagonalization functor and the inclusion of the full subcategory
\<open>\<F>\<^sub>SC\<close> underlie an equivalence of categories. The arrows @{term "mkarr (DOM a\<^bold>\<down>)"},
determined by reductions of canonical representatives, are the components of a
natural isomorphism.
\<close>
interpretation S: full_inclusion_functor comp \<open>\<lambda>f. ide f \<and> Diag (DOM f)\<close> ..
interpretation DoS: composite_functor \<F>\<^sub>SC.comp comp \<F>\<^sub>SC.comp \<F>\<^sub>SC.map D
..
interpretation SoD: composite_functor comp \<F>\<^sub>SC.comp comp D \<F>\<^sub>SC.map ..
interpretation \<nu>: transformation_by_components
comp comp map SoD.map \<open>\<lambda>a. mkarr (DOM a\<^bold>\<down>)\<close>
proof
fix a
assume a: "ide a"
show "\<guillemotleft>mkarr (DOM a\<^bold>\<down>) : map a \<rightarrow> SoD.map a\<guillemotright>"
proof -
have "\<guillemotleft>mkarr (DOM a\<^bold>\<down>) : a \<rightarrow> mkarr \<^bold>\<lfloor>DOM a\<^bold>\<rfloor>\<guillemotright>"
using a Arr_implies_Ide_Dom red_in_Hom dom_char [of a] by auto
moreover have "map a = a"
using a map_simp by simp
moreover have "SoD.map a = mkarr \<^bold>\<lfloor>DOM a\<^bold>\<rfloor>"
using a D.preserves_ide \<F>\<^sub>SC.ideD \<F>\<^sub>SC.map_simp D_def Ide_Diagonalize_rep_ide
Ide_in_Hom Diagonalize_in_Hom
by force
ultimately show ?thesis by simp
qed
next
fix f
assume f: "arr f"
show "mkarr (DOM (cod f)\<^bold>\<down>) \<cdot> map f = SoD.map f \<cdot> mkarr (DOM (dom f)\<^bold>\<down>)"
proof -
have "SoD.map f \<cdot> mkarr (DOM (dom f)\<^bold>\<down>) = mkarr \<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<cdot> mkarr (DOM f\<^bold>\<down>)"
using f DOM_dom D.preserves_arr \<F>\<^sub>SC.map_simp D_def by simp
also have "... = mkarr (\<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<cdot> DOM f\<^bold>\<down>)"
using f Diagonalize_in_Hom red_in_Hom comp_mkarr Arr_implies_Ide_Dom
by simp
also have "... = mkarr (COD f\<^bold>\<down> \<^bold>\<cdot> rep f)"
proof (intro mkarr_eqI)
show "Par (\<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<cdot> DOM f\<^bold>\<down>) (COD f\<^bold>\<down> \<^bold>\<cdot> rep f)"
using f Diagonalize_in_Hom red_in_Hom Arr_implies_Ide_Dom Arr_implies_Ide_Cod
by simp
show "\<^bold>\<lfloor>\<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<cdot> DOM f\<^bold>\<down>\<^bold>\<rfloor> = \<^bold>\<lfloor>COD f\<^bold>\<down> \<^bold>\<cdot> rep f\<^bold>\<rfloor>"
proof -
have "\<^bold>\<lfloor>\<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<cdot> DOM f\<^bold>\<down>\<^bold>\<rfloor> = \<^bold>\<lfloor>rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>DOM f\<^bold>\<down>\<^bold>\<rfloor>"
using f by simp
also have "... = \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using f Arr_implies_Ide_Dom Can_red Ide_Diagonalize_Can [of "DOM f\<^bold>\<down>"]
Diag_Diagonalize CompDiag_Diag_Ide
by force
also have "... = \<^bold>\<lfloor>COD f\<^bold>\<down>\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<cdot>\<^bold>\<rfloor> \<^bold>\<lfloor>rep f\<^bold>\<rfloor>"
using f Arr_implies_Ide_Cod Can_red Ide_Diagonalize_Can [of "COD f\<^bold>\<down>"]
Diag_Diagonalize CompDiag_Diag_Ide
by force
also have "... = \<^bold>\<lfloor>COD f\<^bold>\<down> \<^bold>\<cdot> rep f\<^bold>\<rfloor>"
by simp
finally show ?thesis by blast
qed
qed
also have "... = mkarr (COD f\<^bold>\<down>) \<cdot> mkarr (rep f)"
using f comp_mkarr rep_in_Hom red_in_Hom Arr_implies_Ide_Cod by blast
also have "... = mkarr (DOM (cod f)\<^bold>\<down>) \<cdot> map f"
using f DOM_cod by simp
finally show ?thesis by blast
qed
qed
interpretation \<nu>: natural_isomorphism comp comp map SoD.map \<nu>.map
apply unfold_locales
using \<nu>.map_simp_ide rep_in_Hom Arr_implies_Ide_Dom Can_red can_mkarr_Can iso_can
by simp
text \<open>
The restriction of the diagonalization functor to the subcategory \<open>\<F>\<^sub>SC\<close>
is the identity.
\<close>
lemma DoS_eq_\<F>\<^sub>SC:
shows "DoS.map = \<F>\<^sub>SC.map"
proof
fix f
have "\<not> \<F>\<^sub>SC.arr f \<Longrightarrow> DoS.map f = \<F>\<^sub>SC.map f"
using DoS.is_extensional \<F>\<^sub>SC.map_def by simp
moreover have "\<F>\<^sub>SC.arr f \<Longrightarrow> DoS.map f = \<F>\<^sub>SC.map f"
using \<F>\<^sub>SC.map_simp strict_arr_char Diagonalize_Diag D_def mkarr_Diagonalize_rep
by simp
ultimately show "DoS.map f = \<F>\<^sub>SC.map f" by blast
qed
interpretation \<mu>: transformation_by_components
\<F>\<^sub>SC.comp \<F>\<^sub>SC.comp DoS.map \<F>\<^sub>SC.map \<open>\<lambda>a. a\<close>
using \<F>\<^sub>SC.ideD \<F>\<^sub>SC.map_simp DoS_eq_\<F>\<^sub>SC \<F>\<^sub>SC.map_simp \<F>\<^sub>SC.comp_cod_arr \<F>\<^sub>SC.comp_arr_dom
apply unfold_locales
by (intro \<F>\<^sub>SC.in_homI) auto
interpretation \<mu>: natural_isomorphism \<F>\<^sub>SC.comp \<F>\<^sub>SC.comp DoS.map \<F>\<^sub>SC.map \<mu>.map
apply unfold_locales using \<mu>.map_simp_ide \<F>\<^sub>SC.ide_is_iso by simp
interpretation equivalence_of_categories \<F>\<^sub>SC.comp comp D \<F>\<^sub>SC.map \<nu>.map \<mu>.map ..
text \<open>
We defined the natural isomorphisms @{term \<mu>} and @{term \<nu>} by giving their
components (\emph{i.e.}~their values at objects). However, it is helpful
in exporting these facts to have simple characterizations of their values
for all arrows.
\<close>
definition \<mu>
where "\<mu> \<equiv> \<lambda>f. if \<F>\<^sub>SC.arr f then f else \<F>\<^sub>SC.null"
definition \<nu>
where "\<nu> \<equiv> \<lambda>f. if arr f then mkarr (COD f\<^bold>\<down>) \<cdot> f else null"
lemma \<mu>_char:
shows "\<mu>.map = \<mu>"
proof (intro NaturalTransformation.eqI)
show "natural_transformation \<F>\<^sub>SC.comp \<F>\<^sub>SC.comp DoS.map \<F>\<^sub>SC.map \<mu>.map" ..
have "natural_transformation \<F>\<^sub>SC.comp \<F>\<^sub>SC.comp \<F>\<^sub>SC.map \<F>\<^sub>SC.map \<F>\<^sub>SC.map"
using DoS.as_nat_trans.natural_transformation_axioms DoS_eq_\<F>\<^sub>SC by simp
moreover have "\<F>\<^sub>SC.map = \<mu>" unfolding \<mu>_def using \<F>\<^sub>SC.map_def by blast
ultimately show "natural_transformation \<F>\<^sub>SC.comp \<F>\<^sub>SC.comp DoS.map \<F>\<^sub>SC.map \<mu>"
using \<F>\<^sub>SC.as_nat_trans.natural_transformation_axioms DoS_eq_\<F>\<^sub>SC by simp
show "\<And>a. \<F>\<^sub>SC.ide a \<Longrightarrow> \<mu>.map a = \<mu> a"
using \<mu>.map_simp_ide \<F>\<^sub>SC.ideD \<mu>_def by simp
qed
lemma \<nu>_char:
shows "\<nu>.map = \<nu>"
unfolding \<nu>.map_def \<nu>_def using map_simp DOM_cod by fastforce
lemma is_equivalent_to_strict_subcategory:
shows "equivalence_of_categories \<F>\<^sub>SC.comp comp D \<F>\<^sub>SC.map \<nu> \<mu>"
proof -
have "equivalence_of_categories \<F>\<^sub>SC.comp comp D \<F>\<^sub>SC.map \<nu>.map \<mu>.map" ..
thus "equivalence_of_categories \<F>\<^sub>SC.comp comp D \<F>\<^sub>SC.map \<nu> \<mu>"
using \<nu>_char \<mu>_char by simp
qed
text \<open>
The inclusion of generators functor from @{term C} to \<open>\<F>C\<close>
corestricts to a functor from @{term C} to \<open>\<F>\<^sub>SC\<close>.
\<close>
interpretation I: "functor" C comp inclusion_of_generators
using inclusion_is_functor by auto
interpretation DoI: composite_functor C comp \<F>\<^sub>SC.comp inclusion_of_generators D ..
lemma DoI_eq_I:
shows "DoI.map = inclusion_of_generators"
proof
fix f
have "\<not> C.arr f \<Longrightarrow> DoI.map f = inclusion_of_generators f"
using DoI.is_extensional I.is_extensional \<F>\<^sub>SC.null_char by blast
moreover have "C.arr f \<Longrightarrow> DoI.map f = inclusion_of_generators f"
proof -
assume f: "C.arr f"
have "DoI.map f = D (inclusion_of_generators f)" using f by simp
also have "... = inclusion_of_generators f"
proof -
have "\<F>\<^sub>SC.arr (inclusion_of_generators f)"
using f arr_mkarr rep_mkarr Par_Arr_norm [of "\<^bold>\<langle>f\<^bold>\<rangle>"] strict_arr_char
inclusion_of_generators_def
by simp
thus ?thesis using f strict_arr_char' by blast
qed
finally show "DoI.map f = inclusion_of_generators f" by blast
qed
ultimately show "DoI.map f = inclusion_of_generators f" by blast
qed
end
text \<open>
Next, we show that the subcategory \<open>\<F>\<^sub>SC\<close> inherits monoidal structure from
the ambient category \<open>\<F>C\<close>, and that this monoidal structure is strict.
\<close>
locale free_strict_monoidal_category =
monoidal_language C +
\<F>C: free_monoidal_category C +
full_subcategory \<F>C.comp "\<lambda>f. \<F>C.ide f \<and> Diag (\<F>C.DOM f)"
for C :: "'c comp"
begin
interpretation D: "functor" \<F>C.comp comp \<F>C.D
using \<F>C.diagonalize_is_functor by auto
notation comp (infixr "\<cdot>\<^sub>S" 55)
definition tensor\<^sub>S (infixr "\<otimes>\<^sub>S" 53)
where "f \<otimes>\<^sub>S g \<equiv> \<F>C.D (\<F>C.tensor f g)"
definition assoc\<^sub>S ("\<a>\<^sub>S[_, _, _]")
where "assoc\<^sub>S a b c \<equiv> a \<otimes>\<^sub>S b \<otimes>\<^sub>S c"
lemma tensor_char:
assumes "arr f" and "arr g"
shows "f \<otimes>\<^sub>S g = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor>)"
unfolding \<F>C.D_def tensor\<^sub>S_def
- using assms arr_char \<F>C.rep_tensor by simp
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C \<F>C.rep_tensor by simp
lemma tensor_in_hom [simp]:
assumes "\<guillemotleft>f : a \<rightarrow> b\<guillemotright>" and "\<guillemotleft>g : c \<rightarrow> d\<guillemotright>"
shows "\<guillemotleft>f \<otimes>\<^sub>S g : a \<otimes>\<^sub>S c \<rightarrow> b \<otimes>\<^sub>S d\<guillemotright>"
unfolding tensor\<^sub>S_def
- using assms D.preserves_hom arr_char in_hom_char
+ using assms D.preserves_hom arr_char\<^sub>S\<^sub>b\<^sub>C in_hom_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) \<F>C.T_simp \<F>C.tensor_in_hom in_homE)
lemma arr_tensor [simp]:
assumes "arr f" and "arr g"
shows "arr (f \<otimes>\<^sub>S g)"
using assms arr_iff_in_hom [of f] arr_iff_in_hom [of g] tensor_in_hom by blast
lemma dom_tensor [simp]:
assumes "arr f" and "arr g"
shows "dom (f \<otimes>\<^sub>S g) = dom f \<otimes>\<^sub>S dom g"
using assms arr_iff_in_hom [of f] arr_iff_in_hom [of g] tensor_in_hom by blast
lemma cod_tensor [simp]:
assumes "arr f" and "arr g"
shows "cod (f \<otimes>\<^sub>S g) = cod f \<otimes>\<^sub>S cod g"
using assms arr_iff_in_hom [of f] arr_iff_in_hom [of g] tensor_in_hom by blast
lemma tensor_preserves_ide:
assumes "ide a" and "ide b"
shows "ide (a \<otimes>\<^sub>S b)"
- using assms tensor\<^sub>S_def D.preserves_ide \<F>C.tensor_preserves_ide ide_char
+ using assms tensor\<^sub>S_def D.preserves_ide \<F>C.tensor_preserves_ide ide_char\<^sub>S\<^sub>b\<^sub>C
by fastforce
lemma tensor_tensor:
assumes "arr f" and "arr g" and "arr h"
shows "(f \<otimes>\<^sub>S g) \<otimes>\<^sub>S h = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
and "f \<otimes>\<^sub>S g \<otimes>\<^sub>S h = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
proof -
show "(f \<otimes>\<^sub>S g) \<otimes>\<^sub>S h = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
proof -
have "(f \<otimes>\<^sub>S g) \<otimes>\<^sub>S h = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep (f \<otimes>\<^sub>S g)\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
using assms Diag_Diagonalize TensorDiag_preserves_Diag Diag_implies_Arr
\<F>C.COD_mkarr \<F>C.DOM_mkarr \<F>C.strict_arr_char tensor_char
by simp
also have
"... = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep (\<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor>))\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor>
\<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
- using assms arr_char tensor_char by simp
+ using assms arr_char\<^sub>S\<^sub>b\<^sub>C tensor_char by simp
also have "... = \<F>C.mkarr (\<^bold>\<lfloor>\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor>\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
using assms \<F>C.rep_mkarr TensorDiag_in_Hom Diag_Diagonalize
- TensorDiag_preserves_Diag arr_char
+ TensorDiag_preserves_Diag arr_char\<^sub>S\<^sub>b\<^sub>C
by force
also have "... = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
- using assms Diag_Diagonalize TensorDiag_preserves_Diag TensorDiag_assoc arr_char
+ using assms Diag_Diagonalize TensorDiag_preserves_Diag TensorDiag_assoc arr_char\<^sub>S\<^sub>b\<^sub>C
by force
finally show ?thesis by blast
qed
show "f \<otimes>\<^sub>S g \<otimes>\<^sub>S h = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>)"
proof -
have "... = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>\<^bold>\<rfloor>)"
- using assms Diag_Diagonalize TensorDiag_preserves_Diag arr_char by force
+ using assms Diag_Diagonalize TensorDiag_preserves_Diag arr_char\<^sub>S\<^sub>b\<^sub>C by force
also have "... = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor>
(\<^bold>\<lfloor>\<F>C.rep (\<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep g\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep h\<^bold>\<rfloor>))\<^bold>\<rfloor>))"
- using assms \<F>C.rep_mkarr TensorDiag_in_Hom Diag_Diagonalize arr_char by force
+ using assms \<F>C.rep_mkarr TensorDiag_in_Hom Diag_Diagonalize arr_char\<^sub>S\<^sub>b\<^sub>C by force
also have "... = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep f\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep (g \<otimes>\<^sub>S h)\<^bold>\<rfloor>)"
using assms tensor_char by simp
also have "... = f \<otimes>\<^sub>S g \<otimes>\<^sub>S h"
using assms Diag_Diagonalize TensorDiag_preserves_Diag Diag_implies_Arr
\<F>C.COD_mkarr \<F>C.DOM_mkarr \<F>C.strict_arr_char tensor_char
by simp
finally show ?thesis by blast
qed
qed
lemma tensor_assoc:
assumes "arr f" and "arr g" and "arr h"
shows "(f \<otimes>\<^sub>S g) \<otimes>\<^sub>S h = f \<otimes>\<^sub>S g \<otimes>\<^sub>S h"
using assms tensor_tensor by presburger
lemma arr_unity:
shows "arr \<I>"
using \<F>C.rep_unity \<F>C.Par_Arr_norm \<F>C.\<I>_agreement \<F>C.strict_arr_char by force
lemma tensor_unity_arr:
assumes "arr f"
shows "\<I> \<otimes>\<^sub>S f = f"
using assms arr_unity tensor_char \<F>C.strict_arr_char \<F>C.mkarr_Diagonalize_rep
by simp
lemma tensor_arr_unity:
assumes "arr f"
shows "f \<otimes>\<^sub>S \<I> = f"
using assms arr_unity tensor_char \<F>C.strict_arr_char \<F>C.mkarr_Diagonalize_rep
by simp
lemma assoc_char:
assumes "ide a" and "ide b" and "ide c"
shows "\<a>\<^sub>S[a, b, c] = \<F>C.mkarr (\<^bold>\<lfloor>\<F>C.rep a\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep b\<^bold>\<rfloor> \<^bold>\<lfloor>\<^bold>\<otimes>\<^bold>\<rfloor> \<^bold>\<lfloor>\<F>C.rep c\<^bold>\<rfloor>)"
using assms tensor_tensor(2) assoc\<^sub>S_def ideD(1) by simp
lemma assoc_in_hom:
assumes "ide a" and "ide b" and "ide c"
shows "\<guillemotleft>\<a>\<^sub>S[a, b, c] : (a \<otimes>\<^sub>S b) \<otimes>\<^sub>S c \<rightarrow> a \<otimes>\<^sub>S b \<otimes>\<^sub>S c\<guillemotright>"
using assms tensor_preserves_ide ideD tensor_assoc assoc\<^sub>S_def
by (metis (no_types, lifting) ide_in_hom)
text \<open>The category \<open>\<F>\<^sub>SC\<close> is a monoidal category.\<close>
interpretation EMC: elementary_monoidal_category comp tensor\<^sub>S \<I> \<open>\<lambda>a. a\<close> \<open>\<lambda>a. a\<close> assoc\<^sub>S
proof
show "ide \<I>"
- using ide_char arr_char \<F>C.rep_mkarr \<F>C.Dom_norm \<F>C.Cod_norm \<F>C.\<I>_agreement
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C \<F>C.rep_mkarr \<F>C.Dom_norm \<F>C.Cod_norm \<F>C.\<I>_agreement
by auto
show "\<And>a. ide a \<Longrightarrow> iso a"
- using ide_char arr_char iso_char by auto
+ using ide_char\<^sub>S\<^sub>b\<^sub>C arr_char\<^sub>S\<^sub>b\<^sub>C iso_char\<^sub>S\<^sub>b\<^sub>C by auto
show "\<And>f a b g c d. \<lbrakk> in_hom a b f; in_hom c d g \<rbrakk> \<Longrightarrow> in_hom (a \<otimes>\<^sub>S c) (b \<otimes>\<^sub>S d) (f \<otimes>\<^sub>S g)"
using tensor_in_hom by blast
show "\<And>a b. \<lbrakk> ide a; ide b \<rbrakk> \<Longrightarrow> ide (a \<otimes>\<^sub>S b)"
using tensor_preserves_ide by blast
show "\<And>a b c. \<lbrakk> ide a; ide b; ide c\<rbrakk> \<Longrightarrow> iso \<a>\<^sub>S[a, b, c]"
using tensor_preserves_ide ide_is_iso assoc\<^sub>S_def by presburger
show "\<And>a b c. \<lbrakk> ide a; ide b; ide c\<rbrakk> \<Longrightarrow> \<guillemotleft>\<a>\<^sub>S[a, b, c] : (a \<otimes>\<^sub>S b) \<otimes>\<^sub>S c \<rightarrow> a \<otimes>\<^sub>S b \<otimes>\<^sub>S c\<guillemotright>"
using assoc_in_hom by blast
show "\<And>a b. \<lbrakk> ide a; ide b \<rbrakk> \<Longrightarrow> (a \<otimes>\<^sub>S b) \<cdot>\<^sub>S \<a>\<^sub>S[a, \<I>, b] = a \<otimes>\<^sub>S b"
using ide_def tensor_unity_arr assoc\<^sub>S_def ideD(1) tensor_preserves_ide comp_ide_self
by simp
show "\<And>f. arr f \<Longrightarrow> cod f \<cdot>\<^sub>S (\<I> \<otimes>\<^sub>S f) = f \<cdot>\<^sub>S dom f"
using tensor_unity_arr comp_arr_dom comp_cod_arr by presburger
show "\<And>f. arr f \<Longrightarrow> cod f \<cdot>\<^sub>S (f \<otimes>\<^sub>S \<I>) = f \<cdot>\<^sub>S dom f"
using tensor_arr_unity comp_arr_dom comp_cod_arr by presburger
next
fix a
assume a: "ide a"
show "\<guillemotleft>a : \<I> \<otimes>\<^sub>S a \<rightarrow> a\<guillemotright>"
using a tensor_unity_arr ide_in_hom [of a] by fast
show "\<guillemotleft>a : a \<otimes>\<^sub>S \<I> \<rightarrow> a\<guillemotright>"
using a tensor_arr_unity ide_in_hom [of a] by fast
next
fix f g f' g'
assume fg: "seq g f"
assume fg': "seq g' f'"
show "(g \<otimes>\<^sub>S g') \<cdot>\<^sub>S (f \<otimes>\<^sub>S f') = g \<cdot>\<^sub>S f \<otimes>\<^sub>S g' \<cdot>\<^sub>S f'"
proof -
have A: "\<F>C.seq g f" and B: "\<F>C.seq g' f'"
- using fg fg' seq_char by auto
+ using fg fg' seq_char\<^sub>S\<^sub>b\<^sub>C by auto
have "(g \<otimes>\<^sub>S g') \<cdot>\<^sub>S (f \<otimes>\<^sub>S f') = \<F>C.D ((g \<otimes> g') \<cdot> (f \<otimes> f'))"
using A B tensor\<^sub>S_def by fastforce
also have "... = \<F>C.D (g \<cdot> f \<otimes> g' \<cdot> f')"
using A B \<F>C.interchange \<F>C.T_simp \<F>C.seqE by metis
also have "... = \<F>C.D (g \<cdot> f) \<otimes>\<^sub>S \<F>C.D (g' \<cdot> f')"
- using A B tensor\<^sub>S_def \<F>C.T_simp \<F>C.seqE \<F>C.diagonalize_tensor arr_char
+ using A B tensor\<^sub>S_def \<F>C.T_simp \<F>C.seqE \<F>C.diagonalize_tensor arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting) D.preserves_reflects_arr)
also have "... = \<F>C.D g \<cdot>\<^sub>S \<F>C.D f \<otimes>\<^sub>S \<F>C.D g' \<cdot>\<^sub>S \<F>C.D f'"
using A B by simp
also have "... = g \<cdot>\<^sub>S f \<otimes>\<^sub>S g' \<cdot>\<^sub>S f'"
using fg fg' \<F>C.diagonalize_strict_arr by (elim seqE, simp)
finally show ?thesis by blast
qed
next
fix f0 f1 f2
assume f0: "arr f0" and f1: "arr f1" and f2: "arr f2"
show "\<a>\<^sub>S[cod f0, cod f1, cod f2] \<cdot>\<^sub>S ((f0 \<otimes>\<^sub>S f1) \<otimes>\<^sub>S f2)
= (f0 \<otimes>\<^sub>S f1 \<otimes>\<^sub>S f2) \<cdot>\<^sub>S \<a>\<^sub>S[dom f0, dom f1, dom f2]"
using f0 f1 f2 assoc\<^sub>S_def tensor_assoc dom_tensor cod_tensor arr_tensor
comp_cod_arr [of "f0 \<otimes>\<^sub>S f1 \<otimes>\<^sub>S f2" "cod f0 \<otimes>\<^sub>S cod f1 \<otimes>\<^sub>S cod f2"]
comp_arr_dom [of "f0 \<otimes>\<^sub>S f1 \<otimes>\<^sub>S f2" "dom f0 \<otimes>\<^sub>S dom f1 \<otimes>\<^sub>S dom f2"]
by presburger
next
fix a b c d
assume a: "ide a" and b: "ide b" and c: "ide c" and d: "ide d"
show "(a \<otimes>\<^sub>S \<a>\<^sub>S[b, c, d]) \<cdot>\<^sub>S \<a>\<^sub>S[a, b \<otimes>\<^sub>S c, d] \<cdot>\<^sub>S (\<a>\<^sub>S[a, b, c] \<otimes>\<^sub>S d)
= \<a>\<^sub>S[a, b, c \<otimes>\<^sub>S d] \<cdot>\<^sub>S \<a>\<^sub>S[a \<otimes>\<^sub>S b, c, d]"
unfolding assoc\<^sub>S_def
using a b c d tensor_assoc tensor_preserves_ide ideD tensor_in_hom
comp_arr_dom [of "a \<otimes>\<^sub>S b \<otimes>\<^sub>S c \<otimes>\<^sub>S d"]
by simp
qed
lemma is_elementary_monoidal_category:
shows "elementary_monoidal_category comp tensor\<^sub>S \<I> (\<lambda>a. a) (\<lambda>a. a) assoc\<^sub>S" ..
abbreviation T\<^sub>F\<^sub>S\<^sub>M\<^sub>C where "T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<equiv> EMC.T"
abbreviation \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C where "\<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<equiv> EMC.\<alpha>"
abbreviation \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C where "\<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<equiv> EMC.\<iota>"
lemma is_monoidal_category:
shows "monoidal_category comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C"
using EMC.induces_monoidal_category by auto
end
sublocale free_strict_monoidal_category \<subseteq>
elementary_monoidal_category comp tensor\<^sub>S \<I> "\<lambda>a. a" "\<lambda>a. a" assoc\<^sub>S
using is_elementary_monoidal_category by auto
sublocale free_strict_monoidal_category \<subseteq> monoidal_category comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
using is_monoidal_category by auto
sublocale free_strict_monoidal_category \<subseteq>
strict_monoidal_category comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
using tensor_preserves_ide lunit_agreement runit_agreement \<alpha>_ide_simp assoc\<^sub>S_def
by unfold_locales auto
context free_strict_monoidal_category
begin
text \<open>
The inclusion of generators functor from @{term C} to \<open>\<F>\<^sub>SC\<close> is the composition
of the inclusion of generators from @{term C} to \<open>\<F>C\<close> and the diagonalization
functor, which projects \<open>\<F>C\<close> to \<open>\<F>\<^sub>SC\<close>. As the diagonalization functor
is the identity map on the image of @{term C}, the composite functor amounts to the
corestriction to \<open>\<F>\<^sub>SC\<close> of the inclusion of generators of \<open>\<F>C\<close>.
\<close>
interpretation D: "functor" \<F>C.comp comp \<F>C.D
using \<F>C.diagonalize_is_functor by auto
interpretation I: composite_functor C \<F>C.comp comp \<F>C.inclusion_of_generators \<F>C.D
proof -
interpret "functor" C \<F>C.comp \<F>C.inclusion_of_generators
using \<F>C.inclusion_is_functor by blast
show "composite_functor C \<F>C.comp comp \<F>C.inclusion_of_generators \<F>C.D" ..
qed
definition inclusion_of_generators
where "inclusion_of_generators \<equiv> \<F>C.inclusion_of_generators"
lemma inclusion_is_functor:
shows "functor C comp inclusion_of_generators"
using \<F>C.DoI_eq_I I.functor_axioms inclusion_of_generators_def
by auto
text \<open>
The diagonalization functor is strict monoidal.
\<close>
interpretation D: strict_monoidal_functor \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha>\<^sub>F\<^sub>M\<^sub>C \<F>C.\<iota>\<^sub>F\<^sub>M\<^sub>C
comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
\<F>C.D
proof
show "\<F>C.D \<F>C.\<iota> = \<iota>"
proof -
have "\<F>C.D \<F>C.\<iota> = \<F>C.mkarr \<^bold>\<lfloor>\<F>C.rep \<F>C.\<iota>\<^bold>\<rfloor>"
unfolding \<F>C.D_def using \<F>C.\<iota>_in_hom by auto
also have "... = \<F>C.mkarr \<^bold>\<lfloor>\<^bold>\<l>\<^bold>[\<^bold>\<parallel>\<^bold>\<I>\<^bold>\<parallel>\<^bold>]\<^bold>\<rfloor>"
using \<F>C.\<iota>_def \<F>C.rep_unity \<F>C.rep_lunit \<F>C.Par_Arr_norm \<F>C.Diagonalize_norm
by auto
also have "... = \<iota>"
using \<F>C.unity\<^sub>F\<^sub>M\<^sub>C_def \<F>C.\<I>_agreement \<iota>_def by simp
finally show ?thesis by blast
qed
show "\<And>f g. \<lbrakk> \<F>C.arr f; \<F>C.arr g \<rbrakk> \<Longrightarrow>
\<F>C.D (\<F>C.tensor f g) = tensor (\<F>C.D f) (\<F>C.D g)"
proof -
fix f g
assume f: "\<F>C.arr f" and g: "\<F>C.arr g"
have fg: "arr (\<F>C.D f) \<and> arr (\<F>C.D g)"
using f g D.preserves_arr by blast
have "\<F>C.D (\<F>C.tensor f g) = f \<otimes>\<^sub>S g"
using tensor\<^sub>S_def by simp
also have "f \<otimes>\<^sub>S g = \<F>C.D (f \<otimes> g)"
using f g tensor\<^sub>S_def by simp
also have "... = \<F>C.D f \<otimes>\<^sub>S \<F>C.D g"
- using f g fg tensor\<^sub>S_def \<F>C.T_simp \<F>C.diagonalize_tensor arr_char
+ using f g fg tensor\<^sub>S_def \<F>C.T_simp \<F>C.diagonalize_tensor arr_char\<^sub>S\<^sub>b\<^sub>C
by (metis (no_types, lifting))
also have "... = tensor (\<F>C.D f) (\<F>C.D g)"
using fg T_simp by simp
finally show "\<F>C.D (\<F>C.tensor f g) = tensor (\<F>C.D f) (\<F>C.D g)"
by blast
qed
show "\<And>a b c. \<lbrakk> \<F>C.ide a; \<F>C.ide b; \<F>C.ide c \<rbrakk> \<Longrightarrow>
\<F>C.D (\<F>C.assoc a b c) = assoc (\<F>C.D a) (\<F>C.D b) (\<F>C.D c)"
proof -
fix a b c
assume a: "\<F>C.ide a" and b: "\<F>C.ide b" and c: "\<F>C.ide c"
have abc: "ide (\<F>C.D a) \<and> ide (\<F>C.D b) \<and> ide (\<F>C.D c)"
using a b c D.preserves_ide by blast
have abc': "\<F>C.ide (\<F>C.D a) \<and> \<F>C.ide (\<F>C.D b) \<and> \<F>C.ide (\<F>C.D c)"
- using a b c D.preserves_ide ide_char by simp
+ using a b c D.preserves_ide ide_char\<^sub>S\<^sub>b\<^sub>C by simp
have 1: "\<And>f g. \<F>C.arr f \<Longrightarrow> \<F>C.arr g \<Longrightarrow> f \<otimes>\<^sub>S g = \<F>C.D (f \<otimes> g)"
using tensor\<^sub>S_def by simp
- have 2: "\<And>f. ide f \<Longrightarrow> \<F>C.ide f" using ide_char by blast
+ have 2: "\<And>f. ide f \<Longrightarrow> \<F>C.ide f" using ide_char\<^sub>S\<^sub>b\<^sub>C by blast
have "assoc (\<F>C.D a) (\<F>C.D b) (\<F>C.D c) = \<F>C.D a \<otimes>\<^sub>S \<F>C.D b \<otimes>\<^sub>S \<F>C.D c"
using abc \<alpha>_ide_simp assoc\<^sub>S_def by simp
also have "... = \<F>C.D a \<otimes>\<^sub>S \<F>C.D (\<F>C.D b \<otimes> \<F>C.D c)"
using abc' 1 by auto
also have "... = \<F>C.D a \<otimes>\<^sub>S \<F>C.D (b \<otimes> c)"
using b c \<F>C.diagonalize_tensor by force
also have "... = \<F>C.D (\<F>C.D a \<otimes> \<F>C.D (b \<otimes> c))"
- using 1 b c abc D.preserves_ide \<F>C.tensor_preserves_ide ide_char
+ using 1 b c abc D.preserves_ide \<F>C.tensor_preserves_ide ide_char\<^sub>S\<^sub>b\<^sub>C
by simp
also have "... = \<F>C.D (a \<otimes> b \<otimes> c)"
using a b c \<F>C.diagonalize_tensor by force
also have "... = \<F>C.D \<a>[a, b, c]"
proof -
have "\<F>C.can \<a>[a, b, c]" using a b c \<F>C.can_assoc by simp
hence "\<F>C.ide (\<F>C.D \<a>[a, b, c])"
using a b c \<F>C.ide_diagonalize_can by simp
moreover have "\<F>C.cod (\<F>C.D \<a>[a, b, c]) = \<F>C.D (a \<otimes> b \<otimes> c)"
using a b c \<F>C.assoc_in_hom D.preserves_hom
- by (metis (no_types, lifting) cod_char in_homE)
+ by (metis (no_types, lifting) cod_char\<^sub>S\<^sub>b\<^sub>C in_homE)
ultimately show ?thesis by simp
qed
also have "... = \<F>C.D (\<F>C.assoc a b c)"
using a b c by simp
finally show "\<F>C.D (\<F>C.assoc a b c) = assoc (\<F>C.D a) (\<F>C.D b) (\<F>C.D c)"
by blast
qed
qed
lemma diagonalize_is_strict_monoidal_functor:
shows "strict_monoidal_functor \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha>\<^sub>F\<^sub>M\<^sub>C \<F>C.\<iota>\<^sub>F\<^sub>M\<^sub>C
comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
\<F>C.D"
..
interpretation \<phi>: natural_isomorphism
\<F>C.CC.comp comp D.T\<^sub>DoFF.map D.FoT\<^sub>C.map D.\<phi>
using D.structure_is_natural_isomorphism by simp
text \<open>
The diagonalization functor is part of a monoidal equivalence between the
free monoidal category and the subcategory @{term "\<F>\<^sub>SC"}.
\<close>
interpretation E: equivalence_of_categories comp \<F>C.comp \<F>C.D map \<F>C.\<nu> \<F>C.\<mu>
using \<F>C.is_equivalent_to_strict_subcategory by auto
interpretation D: monoidal_functor \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha>\<^sub>F\<^sub>M\<^sub>C \<F>C.\<iota>\<^sub>F\<^sub>M\<^sub>C
comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
\<F>C.D D.\<phi>
using D.monoidal_functor_axioms by metis
interpretation equivalence_of_monoidal_categories comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
\<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha>\<^sub>F\<^sub>M\<^sub>C \<F>C.\<iota>\<^sub>F\<^sub>M\<^sub>C
\<F>C.D D.\<phi> \<I>
map \<F>C.\<nu> \<F>C.\<mu>
..
text \<open>
The category @{term "\<F>C"} is monoidally equivalent to its subcategory @{term "\<F>\<^sub>SC"}.
\<close>
theorem monoidally_equivalent_to_free_monoidal_category:
shows "equivalence_of_monoidal_categories comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha>\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<iota>\<^sub>F\<^sub>S\<^sub>M\<^sub>C
\<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha>\<^sub>F\<^sub>M\<^sub>C \<F>C.\<iota>\<^sub>F\<^sub>M\<^sub>C
\<F>C.D D.\<phi>
map \<F>C.\<nu> \<F>C.\<mu>"
..
end
text \<open>
We next show that the evaluation functor induced on the free monoidal category
generated by @{term C} by a functor @{term V} from @{term C} to a strict monoidal
category @{term D} restricts to a strict monoidal functor on the subcategory @{term "\<F>\<^sub>SC"}.
\<close>
locale strict_evaluation_functor =
D: strict_monoidal_category D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D +
evaluation_map C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V +
\<F>C: free_monoidal_category C +
E: evaluation_functor C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V +
\<F>\<^sub>SC: free_strict_monoidal_category C
for C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and T\<^sub>D :: "'d * 'd \<Rightarrow> 'd"
and \<alpha>\<^sub>D :: "'d * 'd * 'd \<Rightarrow> 'd"
and \<iota>\<^sub>D :: "'d"
and V :: "'c \<Rightarrow> 'd"
begin
notation \<F>C.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
notation \<F>\<^sub>SC.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
(* TODO: This is just the restriction of the evaluation functor to a subcategory.
It would be useful to define a restriction_of_functor locale that does this in general
and gives the lemma that it yields a functor. *)
definition map
where "map \<equiv> \<lambda>f. if \<F>\<^sub>SC.arr f then E.map f else D.null"
interpretation "functor" \<F>\<^sub>SC.comp D map
unfolding map_def
apply unfold_locales
apply simp
- using \<F>\<^sub>SC.arr_char E.preserves_arr
+ using \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C E.preserves_arr
apply simp
- using \<F>\<^sub>SC.arr_char \<F>\<^sub>SC.dom_char E.preserves_dom
+ using \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.dom_char\<^sub>S\<^sub>b\<^sub>C E.preserves_dom
apply simp
- using \<F>\<^sub>SC.arr_char \<F>\<^sub>SC.cod_char E.preserves_cod
+ using \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.cod_char\<^sub>S\<^sub>b\<^sub>C E.preserves_cod
apply simp
- using \<F>\<^sub>SC.arr_char \<F>\<^sub>SC.dom_char \<F>\<^sub>SC.cod_char \<F>\<^sub>SC.comp_char E.preserves_comp
+ using \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.dom_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.cod_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.comp_char E.preserves_comp
by (elim \<F>\<^sub>SC.seqE, auto)
lemma is_functor:
shows "functor \<F>\<^sub>SC.comp D map" ..
text \<open>
Every canonical arrow is an equivalence class of canonical terms.
The evaluations in \<open>D\<close> of all such terms are identities,
due to the strictness of \<open>D\<close>.
\<close>
lemma ide_eval_Can:
shows "Can t \<Longrightarrow> D.ide \<lbrace>t\<rbrace>"
proof (induct t)
show "\<And>x. Can \<^bold>\<langle>x\<^bold>\<rangle> \<Longrightarrow> D.ide \<lbrace>\<^bold>\<langle>x\<^bold>\<rangle>\<rbrace>" by simp
show "Can \<^bold>\<I> \<Longrightarrow> D.ide \<lbrace>\<^bold>\<I>\<rbrace>" by simp
show "\<And>t1 t2. \<lbrakk> Can t1 \<Longrightarrow> D.ide \<lbrace>t1\<rbrace>; Can t2 \<Longrightarrow> D.ide \<lbrace>t2\<rbrace>; Can (t1 \<^bold>\<otimes> t2) \<rbrakk> \<Longrightarrow>
D.ide \<lbrace>t1 \<^bold>\<otimes> t2\<rbrace>"
by simp
show "\<And>t1 t2. \<lbrakk> Can t1 \<Longrightarrow> D.ide \<lbrace>t1\<rbrace>; Can t2 \<Longrightarrow> D.ide \<lbrace>t2\<rbrace>; Can (t1 \<^bold>\<cdot> t2) \<rbrakk> \<Longrightarrow>
D.ide \<lbrace>t1 \<^bold>\<cdot> t2\<rbrace>"
proof -
fix t1 t2
assume t1: "Can t1 \<Longrightarrow> D.ide \<lbrace>t1\<rbrace>"
and t2: "Can t2 \<Longrightarrow> D.ide \<lbrace>t2\<rbrace>"
and t12: "Can (t1 \<^bold>\<cdot> t2)"
show "D.ide \<lbrace>t1 \<^bold>\<cdot> t2\<rbrace>"
using t1 t2 t12 Can_implies_Arr eval_in_hom [of t1] eval_in_hom [of t2] D.comp_ide_arr
by fastforce
qed
show "\<And>t. (Can t \<Longrightarrow> D.ide \<lbrace>t\<rbrace>) \<Longrightarrow> Can \<^bold>\<l>\<^bold>[t\<^bold>] \<Longrightarrow> D.ide \<lbrace>\<^bold>\<l>\<^bold>[t\<^bold>]\<rbrace>"
using D.strict_lunit by simp
show "\<And>t. (Can t \<Longrightarrow> D.ide \<lbrace>t\<rbrace>) \<Longrightarrow> Can \<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow> D.ide \<lbrace>\<^bold>\<l>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>"
using D.strict_lunit by simp
show "\<And>t. (Can t \<Longrightarrow> D.ide \<lbrace>t\<rbrace>) \<Longrightarrow> Can \<^bold>\<r>\<^bold>[t\<^bold>] \<Longrightarrow> D.ide \<lbrace>\<^bold>\<r>\<^bold>[t\<^bold>]\<rbrace>"
using D.strict_runit by simp
show "\<And>t. (Can t \<Longrightarrow> D.ide \<lbrace>t\<rbrace>) \<Longrightarrow> Can \<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \<Longrightarrow> D.ide \<lbrace>\<^bold>\<r>\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<rbrace>"
using D.strict_runit by simp
fix t1 t2 t3
assume t1: "Can t1 \<Longrightarrow> D.ide \<lbrace>t1\<rbrace>"
and t2: "Can t2 \<Longrightarrow> D.ide \<lbrace>t2\<rbrace>"
and t3: "Can t3 \<Longrightarrow> D.ide \<lbrace>t3\<rbrace>"
show "Can \<^bold>\<a>\<^bold>[t1, t2, t3\<^bold>] \<Longrightarrow> D.ide \<lbrace>\<^bold>\<a>\<^bold>[t1, t2, t3\<^bold>]\<rbrace>"
proof -
assume "Can \<^bold>\<a>\<^bold>[t1, t2, t3\<^bold>]"
hence t123: "D.ide \<lbrace>t1\<rbrace> \<and> D.ide \<lbrace>t2\<rbrace> \<and> D.ide \<lbrace>t3\<rbrace>"
using t1 t2 t3 by simp
have "\<lbrace>\<^bold>\<a>\<^bold>[t1, t2, t3\<^bold>]\<rbrace> = \<lbrace>t1\<rbrace> \<otimes>\<^sub>D \<lbrace>t2\<rbrace> \<otimes>\<^sub>D \<lbrace>t3\<rbrace>"
using t123 D.strict_assoc D.assoc_in_hom [of "\<lbrace>t1\<rbrace>" "\<lbrace>t2\<rbrace>" "\<lbrace>t3\<rbrace>"] apply simp
by (elim D.in_homE, simp)
thus ?thesis using t123 by simp
qed
show "Can \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>] \<Longrightarrow> D.ide \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]\<rbrace>"
proof -
assume "Can \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]"
hence t123: "Can t1 \<and> Can t2 \<and> Can t3 \<and> D.ide \<lbrace>t1\<rbrace> \<and> D.ide \<lbrace>t2\<rbrace> \<and> D.ide \<lbrace>t3\<rbrace>"
using t1 t2 t3 by simp
have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]\<rbrace>
= D.inv \<a>\<^sub>D[D.cod \<lbrace>t1\<rbrace>, D.cod \<lbrace>t2\<rbrace>, D.cod \<lbrace>t3\<rbrace>] \<cdot>\<^sub>D (\<lbrace>t1\<rbrace> \<otimes>\<^sub>D \<lbrace>t2\<rbrace> \<otimes>\<^sub>D \<lbrace>t3\<rbrace>)"
using t123 eval_Assoc' [of t1 t2 t3] Can_implies_Arr by simp
also have "... = \<lbrace>t1\<rbrace> \<otimes>\<^sub>D \<lbrace>t2\<rbrace> \<otimes>\<^sub>D \<lbrace>t3\<rbrace>"
proof -
have "D.dom \<a>\<^sub>D[\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>] = \<lbrace>t1\<rbrace> \<otimes>\<^sub>D \<lbrace>t2\<rbrace> \<otimes>\<^sub>D \<lbrace>t3\<rbrace>"
proof -
have "D.dom \<a>\<^sub>D[\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>] = D.cod \<a>\<^sub>D[\<lbrace>t1\<rbrace>, \<lbrace>t2\<rbrace>, \<lbrace>t3\<rbrace>]"
using t123 D.strict_assoc by simp
also have "... = \<lbrace>t1\<rbrace> \<otimes>\<^sub>D \<lbrace>t2\<rbrace> \<otimes>\<^sub>D \<lbrace>t3\<rbrace>"
using t123 by simp
finally show ?thesis by blast
qed
thus ?thesis
using t123 D.strict_assoc D.comp_arr_dom by auto
qed
finally have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]\<rbrace> = \<lbrace>t1\<rbrace> \<otimes>\<^sub>D \<lbrace>t2\<rbrace> \<otimes>\<^sub>D \<lbrace>t3\<rbrace>" by blast
thus ?thesis using t123 by auto
qed
qed
lemma ide_eval_can:
assumes "\<F>C.can f"
shows "D.ide (E.map f)"
proof -
have "f = \<F>C.mkarr (\<F>C.rep f)"
using assms \<F>C.can_implies_arr \<F>C.mkarr_rep by blast
moreover have 1: "Can (\<F>C.rep f)"
using assms \<F>C.Can_rep_can by simp
moreover have "D.ide \<lbrace>\<F>C.rep f\<rbrace>"
using assms ide_eval_Can by (simp add: 1)
ultimately show ?thesis using assms \<F>C.can_implies_arr E.map_def by force
qed
text \<open>
Diagonalization transports formal arrows naturally along reductions,
which are canonical terms and therefore evaluate to identities of \<open>D\<close>.
It follows that the evaluation in \<open>D\<close> of a formal arrow is equal to the
evaluation of its diagonalization.
\<close>
lemma map_diagonalize:
assumes f: "\<F>C.arr f"
shows "E.map (\<F>C.D f) = E.map f"
proof -
interpret EQ: equivalence_of_categories
\<F>\<^sub>SC.comp \<F>C.comp \<F>C.D \<F>\<^sub>SC.map \<F>C.\<nu> \<F>C.\<mu>
using \<F>C.is_equivalent_to_strict_subcategory by auto
have 1: "\<F>C.seq (\<F>\<^sub>SC.map (\<F>C.D f)) (\<F>C.\<nu> (\<F>C.dom f))"
proof
show "\<guillemotleft>\<F>C.\<nu> (\<F>C.dom f) : \<F>C.dom f \<rightarrow> \<F>C.D (\<F>C.dom f)\<guillemotright>"
using f \<F>\<^sub>SC.map_simp EQ.F.preserves_arr
by (intro \<F>C.in_homI, simp_all)
show "\<guillemotleft>\<F>\<^sub>SC.map (\<F>C.D f) : \<F>C.D (\<F>C.dom f) \<rightarrow> \<F>C.cod (\<F>C.D f)\<guillemotright>"
by (metis (no_types, lifting) EQ.F.preserves_dom EQ.F.preserves_reflects_arr
- \<F>\<^sub>SC.arr_iff_in_hom \<F>\<^sub>SC.cod_simp \<F>\<^sub>SC.in_hom_char \<F>\<^sub>SC.map_simp f)
+ \<F>\<^sub>SC.arr_iff_in_hom \<F>\<^sub>SC.cod_simp \<F>\<^sub>SC.in_hom_char\<^sub>S\<^sub>b\<^sub>C \<F>\<^sub>SC.map_simp f)
qed
have "E.map (\<F>C.\<nu> (\<F>C.cod f)) \<cdot>\<^sub>D E.map f =
E.map (\<F>C.D f) \<cdot>\<^sub>D E.map (\<F>C.\<nu> (\<F>C.dom f))"
proof -
have "E.map (\<F>C.\<nu> (\<F>C.cod f)) \<cdot>\<^sub>D E.map f = E.map (\<F>C.\<nu> (\<F>C.cod f) \<cdot> f)"
using f by simp
also have "... = E.map (\<F>C.D f \<cdot> \<F>C.\<nu> (\<F>C.dom f))"
using f EQ.\<eta>.naturality \<F>\<^sub>SC.map_simp EQ.F.preserves_arr by simp
also have "... = E.map (\<F>\<^sub>SC.map (\<F>C.D f)) \<cdot>\<^sub>D E.map (\<F>C.\<nu> (\<F>C.dom f))"
using f 1 E.as_nat_trans.preserves_comp_2 EQ.F.preserves_arr \<F>\<^sub>SC.map_simp
by (metis (no_types, lifting))
also have "... = E.map (\<F>C.D f) \<cdot>\<^sub>D E.map (\<F>C.\<nu> (\<F>C.dom f))"
using f EQ.F.preserves_arr \<F>\<^sub>SC.map_simp by simp
finally show ?thesis by blast
qed
moreover have "\<And>a. \<F>C.ide a \<Longrightarrow> D.ide (E.map (\<F>C.\<nu> a))"
using \<F>C.\<nu>_def \<F>C.Arr_rep Arr_implies_Ide_Cod Can_red \<F>C.can_mkarr_Can
ide_eval_can
by (metis (no_types, lifting) EQ.\<eta>.preserves_reflects_arr \<F>C.seqE
\<F>C.comp_preserves_can \<F>C.ideD(1) \<F>C.ide_implies_can)
moreover have "D.cod (E.map f) = D.dom (E.map (\<F>C.\<nu> (\<F>C.cod f)))"
using f E.preserves_hom EQ.\<eta>.preserves_hom by simp
moreover have "D.dom (E.map (\<F>C.D f)) = D.cod (E.map (\<F>C.\<nu> (\<F>C.dom f)))"
using f 1 E.preserves_seq EQ.F.preserves_arr \<F>\<^sub>SC.map_simp by auto
ultimately show ?thesis
using f D.comp_arr_dom D.ideD D.arr_dom_iff_arr E.as_nat_trans.is_natural_2
by (metis E.preserves_cod \<F>C.ide_cod \<F>C.ide_dom)
qed
lemma strictly_preserves_tensor:
assumes "\<F>\<^sub>SC.arr f" and "\<F>\<^sub>SC.arr g"
shows "map (\<F>\<^sub>SC.tensor f g) = map f \<otimes>\<^sub>D map g"
proof -
have 1: "\<F>C.arr (f \<otimes> g)"
- using assms \<F>\<^sub>SC.arr_char \<F>C.tensor_in_hom by auto
+ using assms \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C \<F>C.tensor_in_hom by auto
have 2: "\<F>\<^sub>SC.arr (\<F>\<^sub>SC.tensor f g)"
using assms \<F>\<^sub>SC.tensor_in_hom [of f g] \<F>\<^sub>SC.T_simp by fastforce
have "map (\<F>\<^sub>SC.tensor f g) = E.map (f \<otimes> g)"
proof -
have "map (\<F>\<^sub>SC.tensor f g) = map (f \<otimes>\<^sub>S g)"
using assms \<F>\<^sub>SC.T_simp by simp
also have "... = map (\<F>C.D (f \<otimes> g))"
- using assms \<F>C.tensor\<^sub>F\<^sub>M\<^sub>C_def \<F>\<^sub>SC.tensor\<^sub>S_def \<F>\<^sub>SC.arr_char by force
+ using assms \<F>C.tensor\<^sub>F\<^sub>M\<^sub>C_def \<F>\<^sub>SC.tensor\<^sub>S_def \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C by force
also have "... = E.map (f \<otimes> g)"
proof -
interpret Diag: "functor" \<F>C.comp \<F>\<^sub>SC.comp \<F>C.D
using \<F>C.diagonalize_is_functor by auto
show ?thesis
using assms 1 map_diagonalize [of "f \<otimes> g"] Diag.preserves_arr map_def by simp
qed
finally show ?thesis by blast
qed
thus ?thesis
- using assms \<F>\<^sub>SC.arr_char E.strictly_preserves_tensor map_def by simp
+ using assms \<F>\<^sub>SC.arr_char\<^sub>S\<^sub>b\<^sub>C E.strictly_preserves_tensor map_def by simp
qed
lemma is_strict_monoidal_functor:
shows "strict_monoidal_functor \<F>\<^sub>SC.comp \<F>\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<F>\<^sub>SC.\<alpha> \<F>\<^sub>SC.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D map"
proof
show "\<And>f g. \<F>\<^sub>SC.arr f \<Longrightarrow> \<F>\<^sub>SC.arr g \<Longrightarrow> map (\<F>\<^sub>SC.tensor f g) = map f \<otimes>\<^sub>D map g"
using strictly_preserves_tensor by fast
show "map \<F>\<^sub>SC.\<iota> = \<iota>\<^sub>D"
using \<F>\<^sub>SC.arr_unity \<F>\<^sub>SC.\<iota>_def map_def E.map_def \<F>C.rep_mkarr E.eval_norm D.strict_unit
by auto
fix a b c
assume a: "\<F>\<^sub>SC.ide a" and b: "\<F>\<^sub>SC.ide b" and c: "\<F>\<^sub>SC.ide c"
show "map (\<F>\<^sub>SC.assoc a b c) = \<a>\<^sub>D[map a, map b, map c]"
proof -
have "map (\<F>\<^sub>SC.assoc a b c) = map a \<otimes>\<^sub>D map b \<otimes>\<^sub>D map c"
using a b c \<F>\<^sub>SC.\<alpha>_def \<F>\<^sub>SC.assoc\<^sub>S_def \<F>\<^sub>SC.arr_tensor \<F>\<^sub>SC.T_simp \<F>\<^sub>SC.ideD(1)
strictly_preserves_tensor \<F>\<^sub>SC.\<alpha>_ide_simp
by presburger
also have "... = \<a>\<^sub>D[map a, map b, map c]"
using a b c D.strict_assoc D.assoc_in_hom [of "map a" "map b" "map c"] by auto
finally show ?thesis by blast
qed
qed
end
sublocale strict_evaluation_functor \<subseteq>
strict_monoidal_functor \<F>\<^sub>SC.comp \<F>\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<F>\<^sub>SC.\<alpha> \<F>\<^sub>SC.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D map
using is_strict_monoidal_functor by auto
locale strict_monoidal_extension_to_free_strict_monoidal_category =
C: category C +
monoidal_language C +
\<F>\<^sub>SC: free_strict_monoidal_category C +
strict_monoidal_extension C \<F>\<^sub>SC.comp \<F>\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<F>\<^sub>SC.\<alpha> \<F>\<^sub>SC.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
\<F>\<^sub>SC.inclusion_of_generators V F
for C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and T\<^sub>D :: "'d * 'd \<Rightarrow> 'd"
and \<alpha>\<^sub>D :: "'d * 'd * 'd \<Rightarrow> 'd"
and \<iota>\<^sub>D :: "'d"
and V :: "'c \<Rightarrow> 'd"
and F :: "'c free_monoidal_category.arr \<Rightarrow> 'd"
sublocale strict_evaluation_functor \<subseteq>
strict_monoidal_extension C \<F>\<^sub>SC.comp \<F>\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<F>\<^sub>SC.\<alpha> \<F>\<^sub>SC.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
\<F>\<^sub>SC.inclusion_of_generators V map
proof -
interpret V: "functor" C \<F>\<^sub>SC.comp \<F>\<^sub>SC.inclusion_of_generators
using \<F>\<^sub>SC.inclusion_is_functor by auto
show "strict_monoidal_extension C \<F>\<^sub>SC.comp \<F>\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<F>\<^sub>SC.\<alpha> \<F>\<^sub>SC.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
\<F>\<^sub>SC.inclusion_of_generators V map"
proof
show "\<forall>f. C.arr f \<longrightarrow> map (\<F>\<^sub>SC.inclusion_of_generators f) = V f"
using V.preserves_arr E.is_extension map_def \<F>\<^sub>SC.inclusion_of_generators_def by simp
qed
qed
context free_strict_monoidal_category
begin
text \<open>
We now have the main result of this section: the evaluation functor on \<open>\<F>\<^sub>SC\<close>
induced by a functor @{term V} from @{term C} to a strict monoidal category @{term D}
is the unique strict monoidal extension of @{term V} to \<open>\<F>\<^sub>SC\<close>.
\<close>
theorem is_free:
assumes "strict_monoidal_category D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D"
and "strict_monoidal_extension_to_free_strict_monoidal_category C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V F"
shows "F = strict_evaluation_functor.map C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V"
proof -
interpret D: strict_monoidal_category D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D
using assms(1) by auto
text \<open>
Let @{term F} be a given extension of V to a strict monoidal functor defined on
\<open>\<F>\<^sub>SC\<close>.
\<close>
interpret F: strict_monoidal_extension_to_free_strict_monoidal_category
C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V F
using assms(2) by auto
text \<open>
Let @{term E\<^sub>S} be the evaluation functor from \<open>\<F>\<^sub>SC\<close> to @{term D}
induced by @{term V}. Then @{term E\<^sub>S} is also a strict monoidal extension of @{term V}.
\<close>
interpret E\<^sub>S: strict_evaluation_functor C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V ..
text \<open>
Let @{term D} be the strict monoidal functor @{term "\<F>C.D"} that projects
\<open>\<F>C\<close> to the subcategory \<open>\<F>\<^sub>SC\<close>.
\<close>
interpret D: "functor" \<F>C.comp comp \<F>C.D
using \<F>C.diagonalize_is_functor by auto
interpret D: strict_monoidal_functor \<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota>
comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \<alpha> \<iota>
\<F>C.D
using diagonalize_is_strict_monoidal_functor by blast
text \<open>
The composite functor \<open>F o D\<close> is also an extension of @{term V}
to a strict monoidal functor on \<open>\<F>C\<close>.
\<close>
interpret FoD: composite_functor \<F>C.comp comp D \<F>C.D F ..
interpret FoD: strict_monoidal_functor
\<F>C.comp \<F>C.T\<^sub>F\<^sub>M\<^sub>C \<F>C.\<alpha> \<F>C.\<iota> D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D \<open>F o \<F>C.D\<close>
using D.strict_monoidal_functor_axioms F.strict_monoidal_functor_axioms
strict_monoidal_functors_compose
by fast
interpret FoD: strict_monoidal_extension_to_free_monoidal_category
C D T\<^sub>D \<alpha>\<^sub>D \<iota>\<^sub>D V FoD.map
proof
show "\<forall>f. C.arr f \<longrightarrow> FoD.map (\<F>C.inclusion_of_generators f) = V f"
proof -
have "\<And>f. C.arr f \<Longrightarrow> FoD.map (\<F>C.inclusion_of_generators f) = V f"
proof -
fix f
assume f: "C.arr f"
have "FoD.map (\<F>C.inclusion_of_generators f)
= F (\<F>C.D (\<F>C.inclusion_of_generators f))"
using f by simp
also have "... = F (inclusion_of_generators f)"
using f \<F>C.strict_arr_char' F.I.preserves_arr inclusion_of_generators_def by simp
also have "... = V f"
using f F.is_extension by simp
finally show "FoD.map (\<F>C.inclusion_of_generators f) = V f"
by blast
qed
thus ?thesis by blast
qed
qed
text \<open>
By the freeness of \<open>\<F>C\<close>, we have that \<open>F o D\<close>
is equal to the evaluation functor @{term "E\<^sub>S.E.map"} induced by @{term V}
on \<open>\<F>C\<close>. Moreover, @{term E\<^sub>S.map} coincides with @{term "E\<^sub>S.E.map"} on
\<open>\<F>\<^sub>SC\<close> and \<open>F o D\<close> coincides with @{term F} on
\<open>\<F>\<^sub>SC\<close>. Therefore, @{term F} coincides with @{term E} on their common
domain \<open>\<F>\<^sub>SC\<close>, showing @{term "F = E\<^sub>S.map"}.
\<close>
have "\<And>f. arr f \<Longrightarrow> F f = E\<^sub>S.map f"
using \<F>C.strict_arr_char' \<F>C.is_free [of D] E\<^sub>S.E.evaluation_functor_axioms
FoD.strict_monoidal_extension_to_free_monoidal_category_axioms E\<^sub>S.map_def
by simp
moreover have "\<And>f. \<not>arr f \<Longrightarrow> F f = E\<^sub>S.map f"
- using F.is_extensional E\<^sub>S.is_extensional arr_char by auto
+ using F.is_extensional E\<^sub>S.is_extensional arr_char\<^sub>S\<^sub>b\<^sub>C by auto
ultimately show "F = E\<^sub>S.map" by blast
qed
end
end

File Metadata

Mime Type
application/octet-stream
Expires
Fri, May 17, 9:22 PM (2 d)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
P_D1aHrGbRjC
Default Alt Text
(6 MB)

Event Timeline