Theory EquivalenceOfBicategories
section "Equivalence of Bicategories"
text ‹
In this section, we define the notion ``equivalence of bicategories'', which generalizes
the notion of equivalence of categories, and we establish various of its properties.
In particular, we show that ``equivalent bicategories'' is an equivalence relation,
and that a pseudofunctor is part of an equivalence of bicategories if and only if it
is an equivalence pseudofunctor (\emph{i.e.}~it is faithful, locally full, locally
essentially surjective, and biessentially surjective on objects).
›
theory EquivalenceOfBicategories
imports InternalAdjunction PseudonaturalTransformation
begin
subsection "Definition of Equivalence of Bicategories"
text ‹
An equivalence of bicategories between bicategories ‹C› and ‹D› consists of pseudofunctors
‹F : D → C› and ‹G : C → D› and pseudonatural equivalences ‹η: I⇩D ≈ GF› and
‹ε: FG ≈ I⇩C›.
›
locale equivalence_of_bicategories =
C: bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
D: bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
F: pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F +
G: pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G +
FG: composite_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G Φ⇩G F Φ⇩F +
GF: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F G Φ⇩G +
I⇩C: identity_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
I⇩D: identity_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
η: pseudonatural_equivalence V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
I⇩D.map I⇩D.cmp GF.map GF.cmp η⇩0 η⇩1 +
ε: pseudonatural_equivalence V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
FG.map FG.cmp I⇩C.map I⇩C.cmp ε⇩0 ε⇩1
for V⇩C :: "'c comp" (infixr ‹⋅⇩C› 55)
and H⇩C :: "'c comp" (infixr ‹⋆⇩C› 53)
and 𝖺⇩C :: "'c ⇒ 'c ⇒ 'c ⇒ 'c" (‹𝖺⇩C[_, _, _]›)
and 𝗂⇩C :: "'c ⇒ 'c" (‹𝗂⇩C[_]›)
and src⇩C :: "'c ⇒ 'c"
and trg⇩C :: "'c ⇒ 'c"
and V⇩D :: "'d comp" (infixr ‹⋅⇩D› 55)
and H⇩D :: "'d comp" (infixr ‹⋆⇩D› 53)
and 𝖺⇩D :: "'d ⇒ 'd ⇒ 'd ⇒ 'd" (‹𝖺⇩D[_, _, _]›)
and 𝗂⇩D :: "'d ⇒ 'd" (‹𝗂⇩D[_]›)
and src⇩D :: "'d ⇒ 'd"
and trg⇩D :: "'d ⇒ 'd"
and F :: "'d ⇒ 'c"
and Φ⇩F :: "'d * 'd ⇒ 'c"
and G :: "'c ⇒ 'd"
and Φ⇩G :: "'c * 'c ⇒ 'd"
and η⇩0 :: "'d ⇒ 'd"
and η⇩1 :: "'d ⇒ 'd"
and ε⇩0 :: "'c ⇒ 'c"
and ε⇩1 :: "'c ⇒ 'c"
begin
notation C.isomorphic (infix ‹≅⇩C› 50)
notation D.isomorphic (infix ‹≅⇩D› 50)
notation C.iso_in_hom (‹«_ : _ ≅⇩C _»›)
notation D.iso_in_hom (‹«_ : _ ≅⇩D _»›)
notation C.some_quasi_inverse (‹_⇧~⇧C› [1000] 1000)
notation D.some_quasi_inverse (‹_⇧~⇧D› [1000] 1000)
interpretation η': converse_pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
I⇩D.map I⇩D.cmp GF.map GF.cmp η⇩0 η⇩1
..
interpretation ε': converse_pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
FG.map FG.cmp I⇩C.map I⇩C.cmp ε⇩0 ε⇩1
..
text ‹
Although it will be some trouble yet to prove that ‹F› is an equivalence pseudofunctor,
it is not as difficult to prove that the composites ‹GF› and ‹FG› are equivalence
pseudofunctors, by virtue of their being pseudonaturally equivalent to identity
pseudofunctors.
›
interpretation GF: equivalence_pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D GF.map GF.cmp
proof -
interpret GF: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
I⇩D.map I⇩D.cmp GF.map GF.cmp η⇩0 η⇩1
..
show "equivalence_pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D GF.map GF.cmp"
using GF.is_equivalence_pseudofunctor by simp
qed
interpretation FG: equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C FG.map FG.cmp
proof -
interpret FG: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
I⇩C.map I⇩C.cmp FG.map FG.cmp ε'.map⇩0 ε'.map⇩1
..
show "equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C FG.map FG.cmp"
using FG.is_equivalence_pseudofunctor by simp
qed
text ‹
It is also easy to establish that ‹F› and ‹G› are faithful.
We will use the facts, that ‹GF› is locally full and ‹G› is faithful,
to prove that ‹F› is locally full.
›
interpretation F: faithful_functor V⇩D V⇩C F
proof
fix μ μ'
assume par: "D.par μ μ'" and eq: "F μ = F μ'"
have "src⇩D μ = src⇩D μ' ∧ trg⇩D μ = trg⇩D μ'"
using par D.src_dom D.trg_cod by (metis D.src_cod D.trg_cod)
hence "η⇩0 (trg⇩D μ) ⋆⇩D μ = η⇩0 (trg⇩D μ) ⋆⇩D μ'"
using par eq η.iso_map⇩1_ide D.comp_arr_inv' D.trg.preserves_dom
D.comp_arr_dom D.comp_assoc
by (metis GF.is_faithful o_apply)
thus "μ = μ'"
using par η.ide_map⇩0_obj η.components_are_equivalences
D.equivalence_cancel_left [of "η⇩0 (trg⇩D μ)" μ μ']
by simp
qed
interpretation G: faithful_functor V⇩C V⇩D G
proof
show "⋀μ μ'. ⟦C.par μ μ'; G μ = G μ'⟧ ⟹ μ = μ'"
proof -
fix μ μ'
assume par: "C.par μ μ'" and eq: "G μ = G μ'"
have "src⇩C μ = src⇩C μ' ∧ trg⇩C μ = trg⇩C μ'"
using par by (metis C.src_cod C.trg_cod)
hence "μ ⋆⇩C ε⇩0 (src⇩C μ) = μ' ⋆⇩C ε⇩0 (src⇩C μ)"
using par eq ε.iso_map⇩1_ide C.comp_inv_arr' C.src.preserves_dom
C.comp_arr_dom C.comp_assoc
by (metis FG.is_faithful o_apply)
thus "μ = μ'"
using par ε.ide_map⇩0_obj ε.components_are_equivalences
C.equivalence_cancel_right [of "ε⇩0 (src⇩C μ)" μ μ']
by simp
qed
qed
text ‹
It is perhaps not so easy to discover a proof that ‹F› is locally essentially surjective.
Here we follow the proof of \<^cite>‹"johnson-yau-2d-categories"›, Lemma 6.2.13, except we
have expressed it in a way that explicitly shows its constructive nature, given that
we have already chosen an extension of each component of ‹η› and ‹ε› to an adjoint
equivalence.
›
abbreviation Φ
where "Φ ψ f f' ≡ 𝗋⇩C[f'] ⋅⇩C
(f' ⋆⇩C ε'.counit (src⇩C f)) ⋅⇩C
𝖺⇩C[f', ε⇩0 (src⇩C f), ε'.map⇩0 (src⇩C f)] ⋅⇩C
(C.inv (ε⇩1 f') ⋆⇩C ε'.map⇩0 (src⇩C f)) ⋅⇩C
𝖺⇩C⇧-⇧1[ε⇩0 (trg⇩C f), F (G f'), ε'.map⇩0 (src⇩C f)] ⋅⇩C
(ε⇩0 (trg⇩C f) ⋆⇩C F ψ ⋆⇩C ε'.map⇩0 (src⇩C f)) ⋅⇩C
(ε⇩0 (trg⇩C f) ⋆⇩C C.inv (ε'.map⇩1 f)) ⋅⇩C
𝖺⇩C[ε⇩0 (trg⇩C f), ε'.map⇩0 (trg⇩C f), f] ⋅⇩C
(C.inv (ε'.counit (trg⇩C f)) ⋆⇩C f) ⋅⇩C
𝗅⇩C⇧-⇧1[f]"
lemma G_reflects_isomorphic:
assumes "C.ide f" and "C.ide f'" and "src⇩C f = src⇩C f'" and "trg⇩C f = trg⇩C f'"
and "«ψ : G f ≅⇩D G f'»"
shows "«Φ ψ f f' : f ≅⇩C f'»"
proof -
let ?a = "src⇩C f" and ?a' = "trg⇩C f"
have f: "«f : ?a →⇩C ?a'» ∧ C.ide f"
using assms by simp
have f': "«f' : ?a →⇩C ?a'» ∧ C.ide f'"
using assms by simp
have ψ: "«ψ : G.map⇩0 ?a →⇩D G.map⇩0 ?a'»"
using assms D.vconn_implies_hpar(1-2)
by (elim D.iso_in_homE) auto
hence Fψ: "«F ψ : F.map⇩0 (G.map⇩0 ?a) →⇩C F.map⇩0 (G.map⇩0 ?a')»"
by auto
show "«Φ ψ f f' : f ≅⇩C f'»"
proof (intro C.comp_iso_in_hom)
show "«𝗅⇩C⇧-⇧1[f] : f ≅⇩C ?a' ⋆⇩C f»"
using f by auto
show "«C.inv (ε'.counit ?a') ⋆⇩C f : ?a' ⋆⇩C f ≅⇩C (ε⇩0 ?a' ⋆⇩C ε'.map⇩0 ?a') ⋆⇩C f»"
using f ε'.counit_in_hom [of ?a'] ε'.iso_counit [of ?a']
by (intro C.hcomp_iso_in_hom) auto
show "«𝖺⇩C[ε⇩0 ?a', ε'.map⇩0 ?a', f] :
(ε⇩0 ?a' ⋆⇩C ε'.map⇩0 ?a') ⋆⇩C f ≅⇩C ε⇩0 ?a' ⋆⇩C ε'.map⇩0 ?a' ⋆⇩C f»"
using f ε'.counit_in_hom [of ?a] ε'.ide_map⇩0_obj
by (intro C.iso_in_homI) auto
show "«ε⇩0 ?a' ⋆⇩C C.inv (ε'.map⇩1 f) :
ε⇩0 ?a' ⋆⇩C ε'.map⇩0 ?a' ⋆⇩C f ≅⇩C ε⇩0 ?a' ⋆⇩C F (G f) ⋆⇩C ε'.map⇩0 ?a»"
using f ε'.map⇩1_in_hom [of f] ε'.iso_map⇩1_ide [of f] C.ide_iso_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "«ε⇩0 ?a' ⋆⇩C F ψ ⋆⇩C ε'.map⇩0 ?a :
ε⇩0 ?a' ⋆⇩C F (G f) ⋆⇩C ε'.map⇩0 ?a ≅⇩C ε⇩0 ?a' ⋆⇩C F (G f') ⋆⇩C ε'.map⇩0 ?a»"
using assms f f' Fψ F.preserves_iso C.ide_iso_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "«𝖺⇩C⇧-⇧1[ε⇩0 ?a', F (G f'), ε'.map⇩0 ?a] :
ε⇩0 ?a' ⋆⇩C F (G f') ⋆⇩C ε'.map⇩0 ?a ≅⇩C (ε⇩0 ?a' ⋆⇩C F (G f')) ⋆⇩C ε'.map⇩0 ?a»"
using assms f' ε.map⇩0_in_hom(1) [of ?a'] ε.ide_map⇩0_obj [of ?a']
ε'.map⇩0_in_hom(1) [of ?a] ε'.ide_map⇩0_obj [of ?a]
C.assoc'_in_hom C.iso_assoc'
by auto
show "«C.inv (ε⇩1 f') ⋆⇩C ε'.map⇩0 ?a :
(ε⇩0 ?a' ⋆⇩C F (G f')) ⋆⇩C ε'.map⇩0 ?a ≅⇩C (f' ⋆⇩C ε⇩0 ?a) ⋆⇩C ε'.map⇩0 ?a»"
using assms f' ε.map⇩1_in_hom ε.iso_map⇩1_ide ε'.map⇩0_in_hom(1) [of ?a]
C.ide_iso_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "«𝖺⇩C[f', ε⇩0 ?a, ε'.map⇩0 ?a] :
(f' ⋆⇩C ε⇩0 ?a) ⋆⇩C ε'.map⇩0 ?a ≅⇩C f' ⋆⇩C ε⇩0 ?a ⋆⇩C ε'.map⇩0 ?a»"
using assms f' ε.map⇩0_in_hom(1) [of ?a] ε'.map⇩0_in_hom(1) [of ?a]
ε.ide_map⇩0_obj ε'.ide_map⇩0_obj C.assoc_in_hom
by auto
show "«f' ⋆⇩C ε'.counit ?a : f' ⋆⇩C ε⇩0 ?a ⋆⇩C ε'.map⇩0 ?a ≅⇩C f' ⋆⇩C ?a»"
using f f' ε'.counit_in_hom
by (intro C.hcomp_iso_in_hom) auto
show "«𝗋⇩C[f'] : f' ⋆⇩C ?a ≅⇩C f'»"
using assms f' by auto
qed
qed
abbreviation Ψ
where "Ψ f b b' ≡ 𝗋⇩D[G (F (η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b))] ⋅⇩D
(G (F (η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b)) ⋆⇩D η'.ε b) ⋅⇩D
𝖺⇩D[G (F (η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b)), η⇩0 b, η'.map⇩0 b] ⋅⇩D
(D.inv (η⇩1 (η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b)) ⋆⇩D η'.map⇩0 b) ⋅⇩D
(𝖺⇩D[η⇩0 b', η'.map⇩0 b', G f ⋆⇩D η⇩0 b] ⋆⇩D η'.map⇩0 b) ⋅⇩D
𝖺⇩D⇧-⇧1[η⇩0 b' ⋆⇩D η'.map⇩0 b', G f ⋆⇩D η⇩0 b, η'.map⇩0 b] ⋅⇩D
((η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D 𝖺⇩D⇧-⇧1[G f, η⇩0 b, η'.map⇩0 b]) ⋅⇩D
(D.inv (η'.counit b') ⋆⇩D G f ⋆⇩D D.inv (η'.counit b)) ⋅⇩D
𝗅⇩D⇧-⇧1[G f ⋆⇩D G.map⇩0 (F.map⇩0 b)] ⋅⇩D
𝗋⇩D⇧-⇧1[G f]"
lemma F_is_locally_essentially_surjective:
assumes "D.obj b" and "D.obj b'" and "«f : F.map⇩0 b →⇩C F.map⇩0 b'»" and "C.ide f"
shows "«Φ (Ψ f b b') f (F (η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b)) :
f ≅⇩C F (η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b)»"
proof -
let ?g = "η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b"
have g_in_hhom: "«?g : b →⇩D b'»"
using assms η.ide_map⇩0_obj η.map⇩0_in_hhom by auto
have ide_g: "D.ide ?g"
using assms g_in_hhom η.ide_map⇩0_obj η'.ide_map⇩0_obj G.preserves_ide by blast
let ?ψ = "Ψ f b b'"
have "«?ψ : G f ≅⇩D G (F ?g)»"
proof (intro D.comp_iso_in_hom)
show "«𝗋⇩D⇧-⇧1[G f] : G f ≅⇩D G f ⋆⇩D G.map⇩0 (F.map⇩0 b)»"
using assms
by (intro D.iso_in_homI) auto
show "«𝗅⇩D⇧-⇧1[G f ⋆⇩D G.map⇩0 (F.map⇩0 b)] :
G f ⋆⇩D G.map⇩0 (F.map⇩0 b)
≅⇩D G.map⇩0 (F.map⇩0 b') ⋆⇩D G f ⋆⇩D G.map⇩0 (F.map⇩0 b)»"
proof -
have "D.ide (G f ⋆⇩D G.map⇩0 (F.map⇩0 b))"
using assms D.ide_hcomp [of "G f" "G.map⇩0 (F.map⇩0 b)"]
by (metis D.hcomp_in_hhomE D.hseqE D.objE F.map⇩0_simps(1) G.map⇩0_simps(1)
G.preserves_ide GF.map⇩0_simp η.map⇩0_simps(3) g_in_hhom)
thus ?thesis
using assms η.ide_map⇩0_obj
by (intro D.iso_in_homI) auto
qed
show "«D.inv (η'.counit b') ⋆⇩D G f ⋆⇩D D.inv (η'.counit b) :
G.map⇩0 (F.map⇩0 b') ⋆⇩D G f ⋆⇩D G.map⇩0 (F.map⇩0 b)
≅⇩D (η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D G f ⋆⇩D (η⇩0 b ⋆⇩D η'.map⇩0 b)»"
using assms η'.iso_counit η'.counit_in_hom(2) D.vconn_implies_hpar(4) D.ide_iso_in_hom
by (intro D.hcomp_iso_in_hom) auto
show "«(η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D 𝖺⇩D⇧-⇧1[G f, η⇩0 b, η'.map⇩0 b] :
(η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D G f ⋆⇩D (η⇩0 b ⋆⇩D η'.map⇩0 b)
≅⇩D (η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D (G f ⋆⇩D η⇩0 b) ⋆⇩D η'.map⇩0 b»"
proof -
have "«𝖺⇩D⇧-⇧1[G f, η⇩0 b, η'.map⇩0 b] :
G f ⋆⇩D (η⇩0 b ⋆⇩D η'.map⇩0 b) ≅⇩D (G f ⋆⇩D η⇩0 b) ⋆⇩D η'.map⇩0 b»"
using assms η.ide_map⇩0_obj η'.ide_map⇩0_obj η.map⇩0_in_hom η'.map⇩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 "«𝖺⇩D⇧-⇧1[η⇩0 b' ⋆⇩D η'.map⇩0 b', G f ⋆⇩D η⇩0 b, η'.map⇩0 b] :
(η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D (G f ⋆⇩D η⇩0 b) ⋆⇩D η'.map⇩0 b
≅⇩D ((η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D (G f ⋆⇩D η⇩0 b)) ⋆⇩D η'.map⇩0 b»"
using assms η.ide_map⇩0_obj η'.ide_map⇩0_obj η.map⇩0_in_hom η'.map⇩0_in_hom
D.assoc'_in_hom D.iso_assoc'
by (intro D.iso_in_homI) auto
show "«𝖺⇩D[η⇩0 b', η'.map⇩0 b', G f ⋆⇩D η⇩0 b] ⋆⇩D η'.map⇩0 b :
((η⇩0 b' ⋆⇩D η'.map⇩0 b') ⋆⇩D (G f ⋆⇩D η⇩0 b)) ⋆⇩D η'.map⇩0 b
≅⇩D (η⇩0 b' ⋆⇩D ?g) ⋆⇩D η'.map⇩0 b»"
using assms η.ide_map⇩0_obj η'.ide_map⇩0_obj η.map⇩0_in_hom η'.map⇩0_in_hom
D.assoc_in_hom D.iso_assoc
by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto
show "«D.inv (η⇩1 ?g) ⋆⇩D η'.map⇩0 b :
(η⇩0 b' ⋆⇩D ?g) ⋆⇩D η'.map⇩0 b ≅⇩D (G (F ?g) ⋆⇩D η⇩0 b) ⋆⇩D η'.map⇩0 b»"
using assms η.map⇩1_in_hom [of ?g] η.iso_map⇩1_ide η'.map⇩0_in_hom g_in_hhom ide_g
by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto
show "«𝖺⇩D[G (F ?g), η⇩0 b, η'.map⇩0 b] :
(G (F ?g) ⋆⇩D η⇩0 b) ⋆⇩D η'.map⇩0 b ≅⇩D G (F ?g) ⋆⇩D η⇩0 b ⋆⇩D η'.map⇩0 b»"
using assms η.ide_map⇩0_obj η'.ide_map⇩0_obj η.map⇩0_in_hom η'.map⇩0_in_hom
D.assoc_in_hom D.iso_assoc
by (intro D.iso_in_homI) auto
show "«G (F ?g) ⋆⇩D η'.counit b :
G (F ?g) ⋆⇩D η⇩0 b ⋆⇩D η'.map⇩0 b ≅⇩D G (F ?g) ⋆⇩D G.map⇩0 (F.map⇩0 b)»"
using assms η'.counit_in_hom D.ide_in_hom(2) ide_g
by (intro D.hcomp_iso_in_hom) auto
show "«𝗋⇩D[G (F ?g)] : G (F ?g) ⋆⇩D G.map⇩0 (F.map⇩0 b) ≅⇩D G (F ?g)»"
using assms ide_g
by (intro D.iso_in_homI) auto
qed
thus "«Φ ?ψ f (F ?g) : f ≅⇩C F ?g»"
using assms ide_g G_reflects_isomorphic [of f "F ?g" ?ψ]
by (metis C.horizontal_homs_axioms D.in_hhomE F.preserves_ide F.preserves_src
F.preserves_trg g_in_hhom horizontal_homs.in_hhomE)
qed
interpretation F: equivalence_pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F
proof
show "⋀b. C.obj b ⟹ ∃a. D.obj a ∧ C.equivalent_objects (F.map⇩0 a) b"
by (metis FG.biessentially_surjective_on_objects FG.map⇩0_simp G.map⇩0_simps(1))
show "⋀g g' μ. ⟦D.ide g; D.ide g'; src⇩D g = src⇩D g'; trg⇩D g = trg⇩D g';
«μ : F g ⇒⇩C F g'»⟧
⟹ ∃ν. «ν : g ⇒⇩D g'» ∧ F ν = μ"
proof -
fix g g' μ
assume g: "D.ide g" and g': "D.ide g'"
assume eq_src: "src⇩D g = src⇩D g'" and eq_trg: "trg⇩D g = trg⇩D g'"
assume μ: "«μ : F g ⇒⇩C F g'»"
obtain ν where ν: "«ν : g ⇒⇩D g'» ∧ G (F ν) = G μ"
using g g' eq_src eq_trg μ GF.locally_full [of g g' "G μ"] by auto
have "F ν = μ"
using μ ν G.is_faithful
by (metis (mono_tags, lifting) C.in_homE F.preserves_hom)
thus "∃ν. «ν : g ⇒⇩D g'» ∧ F ν = μ"
using ν by auto
qed
show "⋀b b' f. ⟦D.obj b; D.obj b'; «f : F.map⇩0 b →⇩C F.map⇩0 b'»; C.ide f⟧
⟹ ∃g. «g : b →⇩D b'» ∧ D.ide g ∧ C.isomorphic (F g) f"
proof -
fix b b' f
assume b: "D.obj b" and b': "D.obj b'" and f: "«f : F.map⇩0 b →⇩C F.map⇩0 b'»"
assume ide_f: "C.ide f"
let ?g = "η'.map⇩0 b' ⋆⇩D G f ⋆⇩D η⇩0 b"
have g_in_hhom: "«?g : b →⇩D b'»"
using b b' f η.ide_map⇩0_obj η.map⇩0_in_hhom by auto
have ide_g: "D.ide ?g"
using b b' f ide_f g_in_hhom η.ide_map⇩0_obj η'.ide_map⇩0_obj G.preserves_ide by blast
have "f ≅⇩C F ?g"
using b b' f ide_f F_is_locally_essentially_surjective C.isomorphic_symmetric
by (meson C.isomorphicI')
thus "∃g. «g : b →⇩D b'» ∧ D.ide g ∧ F g ≅⇩C f"
using g_in_hhom ide_g C.isomorphic_symmetric C.isomorphic_def by blast
qed
qed
lemma equivalence_pseudofunctor_left:
shows "equivalence_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F"
..
end
subsection "Equivalences Respect Pseudonatural Equivalence"
text ‹
In this section we prove that, in an equivalence of bicategories, either pseudofunctor
may be replaced by a pseudonaturally equivalent one and still obtain an equivalence of
bicategories.
›
locale equivalence_of_bicategories_and_pseudonatural_equivalence_left =
C: bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
D: bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
E: equivalence_of_bicategories +
τ: pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F F' Φ⇩F' τ⇩0 τ⇩1
for F'
and Φ⇩F'
and τ⇩0
and τ⇩1
begin
interpretation GF': composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F' Φ⇩F' G Φ⇩G
..
interpretation F'G: composite_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G Φ⇩G F' Φ⇩F'
..
interpretation τ': converse_pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F F' Φ⇩F' τ⇩0 τ⇩1
..
interpretation τ'oG: pseudonatural_equivalence_whisker_right V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F' Φ⇩F' F Φ⇩F G Φ⇩G τ'.map⇩0 τ'.map⇩1
..
interpretation Goτ: pseudonatural_equivalence_whisker_left V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ⇩F F' Φ⇩F' G Φ⇩G τ⇩0 τ⇩1
..
sublocale unit: composite_pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
E.I⇩D.map E.I⇩D.cmp E.GF.map E.GF.cmp GF'.map GF'.cmp
η⇩0 η⇩1 Goτ.map⇩0 Goτ.map⇩1
..
sublocale counit: composite_pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F'G.map F'G.cmp E.FG.map E.FG.cmp E.I⇩C.map E.I⇩C.cmp
τ'oG.map⇩0 τ'oG.map⇩1 ε⇩0 ε⇩1
..
definition unit⇩0
where "unit⇩0 ≡ unit.map⇩0"
definition unit⇩1
where "unit⇩1 ≡ unit.map⇩1"
definition counit⇩0
where "counit⇩0 ≡ counit.map⇩0"
definition counit⇩1
where "counit⇩1 ≡ counit.map⇩1"
sublocale equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F' Φ⇩F' G Φ⇩G
unit⇩0 unit⇩1 counit⇩0 counit⇩1
unfolding unit⇩0_def unit⇩1_def counit⇩0_def counit⇩1_def
..
lemma induces_equivalence_of_bicategories:
shows "equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F' Φ⇩F' G Φ⇩G
unit⇩0 unit⇩1 counit⇩0 counit⇩1"
..
end
locale equivalence_of_bicategories_and_pseudonatural_equivalence_right =
C: bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
D: bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
E: equivalence_of_bicategories +
τ: pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G G' Φ⇩G' τ⇩0 τ⇩1
for G'
and Φ⇩G'
and τ⇩0
and τ⇩1
begin
interpretation G'F: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F G' Φ⇩G'
..
interpretation FG': composite_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G' Φ⇩G' F Φ⇩F
..
interpretation τ': converse_pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G G' Φ⇩G' τ⇩0 τ⇩1
..
interpretation τoF: pseudonatural_equivalence_whisker_right V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
G Φ⇩G G' Φ⇩G' F Φ⇩F τ⇩0 τ⇩1
..
interpretation Foτ': pseudonatural_equivalence_whisker_left
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
G' Φ⇩G' G Φ⇩G F Φ⇩F τ'.map⇩0 τ'.map⇩1
..
sublocale unit: composite_pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
E.I⇩D.map E.I⇩D.cmp E.GF.map E.GF.cmp G'F.map G'F.cmp
η⇩0 η⇩1 τoF.map⇩0 τoF.map⇩1
..
sublocale counit: composite_pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
FG'.map FG'.cmp E.FG.map E.FG.cmp E.I⇩C.map E.I⇩C.cmp
Foτ'.map⇩0 Foτ'.map⇩1 ε⇩0 ε⇩1
..
definition unit⇩0
where "unit⇩0 ≡ unit.map⇩0"
definition unit⇩1
where "unit⇩1 ≡ unit.map⇩1"
definition counit⇩0
where "counit⇩0 ≡ counit.map⇩0"
definition counit⇩1
where "counit⇩1 ≡ counit.map⇩1"
sublocale equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F G' Φ⇩G'
unit⇩0 unit⇩1 counit⇩0 counit⇩1
unfolding unit⇩0_def unit⇩1_def counit⇩0_def counit⇩1_def
..
lemma induces_equivalence_of_bicategories:
shows "equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F G' Φ⇩G'
unit⇩0 unit⇩1 counit⇩0 counit⇩1"
..
end
subsection "Converse of an Equivalence"
text ‹
Equivalence of bicategories is a symmetric notion, in the sense that from an equivalence
of bicategories from ‹C› to ‹D› we may obtain an equivalence of bicategories from
‹D› to ‹C›. The converse equivalence is obtained by interchanging the pseudofunctors
‹F› and ‹G› and replacing the pseudonatural equivalences ‹η› and ‹ε› by converse
equivalences. Essentially all the work goes into proving that pseudonatural equivalences
have pseudonatural converses, which we have already done.
›
locale converse_equivalence_of_bicategories =
C: bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
D: bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
I⇩C: identity_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
I⇩D: identity_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
E: equivalence_of_bicategories
begin
sublocale counit: converse_pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
I⇩D.map I⇩D.cmp E.GF.map E.GF.cmp η⇩0 η⇩1
..
sublocale unit: converse_pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
E.FG.map E.FG.cmp I⇩C.map I⇩C.cmp ε⇩0 ε⇩1
..
definition unit⇩0
where "unit⇩0 ≡ unit.map⇩0"
definition unit⇩1
where "unit⇩1 ≡ unit.map⇩1"
definition counit⇩0
where "counit⇩0 ≡ counit.map⇩0"
definition counit⇩1
where "counit⇩1 ≡ counit.map⇩1"
sublocale equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G Φ⇩G F Φ⇩F
unit⇩0 unit⇩1 counit⇩0 counit⇩1
unfolding unit⇩0_def unit⇩1_def counit⇩0_def counit⇩1_def
..
lemma is_equivalence_of_bicategories:
shows "equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G Φ⇩G F Φ⇩F
unit⇩0 unit⇩1 counit⇩0 counit⇩1"
..
end
subsection "Composition of Equivalences"
text ‹
An equivalence of bicategories from ‹B› to ‹C› and an equivalence of bicategories
from ‹C› to ‹D› may be composed to obtain an equivalence of bicategories
from ‹B› to ‹D›.
›
locale composite_equivalence_of_bicategories =
B: bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B +
C: bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
D: bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
I⇩B: identity_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B +
I⇩C: identity_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
I⇩D: identity_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
F_: pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B F Φ⇩F +
G_: pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G Φ⇩G +
H: pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C H Φ⇩H +
K: pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D K Φ⇩K +
FG: equivalence_of_bicategories
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F G Φ⇩G ρ⇩0 ρ⇩1 σ⇩0 σ⇩1 +
HK: equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D H Φ⇩H K Φ⇩K ζ⇩0 ζ⇩1 ξ⇩0 ξ⇩1
for V⇩B :: "'b comp" (infixr ‹⋅⇩B› 55)
and H⇩B :: "'b comp" (infixr ‹⋆⇩B› 53)
and 𝖺⇩B :: "'b ⇒ 'b ⇒ 'b ⇒ 'b" (‹𝖺⇩B[_, _, _]›)
and 𝗂⇩B :: "'b ⇒ 'b" (‹𝗂⇩B[_]›)
and src⇩B :: "'b ⇒ 'b"
and trg⇩B :: "'b ⇒ 'b"
and V⇩C :: "'c comp" (infixr ‹⋅⇩C› 55)
and H⇩C :: "'c comp" (infixr ‹⋆⇩C› 53)
and 𝖺⇩C :: "'c ⇒ 'c ⇒ 'c ⇒ 'c" (‹𝖺⇩C[_, _, _]›)
and 𝗂⇩C :: "'c ⇒ 'c" (‹𝗂⇩C[_]›)
and src⇩C :: "'c ⇒ 'c"
and trg⇩C :: "'c ⇒ 'c"
and V⇩D :: "'d comp" (infixr ‹⋅⇩D› 55)
and H⇩D :: "'d comp" (infixr ‹⋆⇩D› 53)
and 𝖺⇩D :: "'d ⇒ 'd ⇒ 'd ⇒ 'd" (‹𝖺⇩D[_, _, _]›)
and 𝗂⇩D :: "'d ⇒ 'd" (‹𝗂⇩D[_]›)
and src⇩D :: "'d ⇒ 'd"
and trg⇩D :: "'d ⇒ 'd"
and F :: "'c ⇒ 'b"
and Φ⇩F :: "'c * 'c ⇒ 'b"
and G :: "'b ⇒ 'c"
and Φ⇩G :: "'b * 'b ⇒ 'c"
and H :: "'d ⇒ 'c"
and Φ⇩H :: "'d * 'd ⇒ 'c"
and K :: "'c ⇒ 'd"
and Φ⇩K :: "'c * 'c ⇒ 'd"
and ρ⇩0 :: "'c ⇒ 'c"
and ρ⇩1 :: "'c ⇒ 'c"
and σ⇩0 :: "'b ⇒ 'b"
and σ⇩1 :: "'b ⇒ 'b"
and ζ⇩0 :: "'d ⇒ 'd"
and ζ⇩1 :: "'d ⇒ 'd"
and ξ⇩0 :: "'c ⇒ 'c"
and ξ⇩1 :: "'c ⇒ 'c"
begin
notation B.𝖺' (‹𝖺⇩B⇧-⇧1[_, _, _]›)
text ‹
At this point we could make the explicit definitions:
\begin{itemize}
\item ‹η⇩0 = K (ρ⇩0 (H.map⇩0 a)) ⋆⇩D ζ⇩0 a›
\item ‹η⇩1 = 𝖺⇩D⇧-⇧1[K (ρ⇩0 (H.map⇩0 (trg⇩D f))), ζ⇩0 (trg⇩D f), f] ⋅⇩D
(K (ρ⇩0 (H.map⇩0 (trg⇩D f))) ⋆⇩D ζ⇩1 f) ⋅⇩D
𝖺⇩D[K (ρ⇩0 (H.map⇩0 (trg⇩D f))), K (H f), ζ⇩0 (src⇩D f)] ⋅⇩D
(D.inv (Φ⇩K (ρ⇩0 (H.map⇩0 (trg⇩D f)), H f)) ⋅⇩D
K (ρ⇩1 (H f)) ⋅⇩D
Φ⇩K (G (F (H f)), ρ⇩0 (H.map⇩0 (src⇩D f))) ⋆⇩D ζ⇩0 (src⇩D f)) ⋅⇩D
𝖺⇩D⇧-⇧1[K (G (F (H f))), K (ρ⇩0 (H.map⇩0 (src⇩D f))), ζ⇩0 (src⇩D f)]›
\item ‹ε⇩0 = σ⇩0 a ⋆⇩B F (ξ⇩0 (G_.map⇩0 a))›
\item ‹ε⇩1 = 𝖺⇩B⇧-⇧1[σ⇩0 (trg⇩B f), F (ξ⇩0 (G_.map⇩0 (trg⇩B f))), F (H (K (G f)))] ⋅⇩B
(σ⇩0 (trg⇩B f) ⋆⇩B
B.inv (Φ⇩F (ξ⇩0 (G_.map⇩0 (trg⇩B f)), H (K (G f)))) ⋅⇩B
F (ξ⇩1 (G f)) ⋅⇩B Φ⇩F (G f, ξ⇩0 (G_.map⇩0 (src⇩B f)))) ⋅⇩B
𝖺⇩B[σ⇩0 (trg⇩B f), F (G f), F (ξ⇩0 (G_.map⇩0 (src⇩B f)))] ⋅⇩B
(σ⇩1 f ⋆⇩B F (ξ⇩0 (G_.map⇩0 (src⇩B f)))) ⋅⇩B
𝖺⇩B⇧-⇧1[f, σ⇩0 (src⇩B f), F (ξ⇩0 (G_.map⇩0 (src⇩B f)))]›
\end{itemize}
but then it is a daunting task to establish the necessary coherence conditions.
It is easier (and more useful) to use general results about composite pseudonatural
equivalences, which are somewhat easier to prove, though long calculations are
still required for those.
›
sublocale FH: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B H Φ⇩H F Φ⇩F
..
sublocale KG: composite_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G K Φ⇩K
..
interpretation IG: composite_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
G Φ⇩G I⇩C.map I⇩C.cmp
..
interpretation IH: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
H Φ⇩H I⇩C.map I⇩C.cmp
..
interpretation HKG: composite_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
G Φ⇩G HK.FG.map HK.FG.cmp
..
interpretation GFH: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
H Φ⇩H FG.GF.map FG.GF.cmp
..
interpretation KGFH: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
GFH.map GFH.cmp K Φ⇩K
..
interpretation FHKG: composite_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
HKG.map HKG.cmp F Φ⇩F
..
interpretation ρoH: pseudonatural_equivalence_whisker_right V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
I⇩C.map I⇩C.cmp FG.GF.map FG.GF.cmp H Φ⇩H ρ⇩0 ρ⇩1
..
interpretation KoρoH: pseudonatural_equivalence_whisker_left V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
H Φ⇩H GFH.map GFH.cmp K Φ⇩K ρoH.map⇩0 ρoH.map⇩1
proof -
interpret KoρoH: pseudonatural_equivalence_whisker_left V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
IH.map IH.cmp GFH.map GFH.cmp K Φ⇩K ρoH.map⇩0 ρoH.map⇩1
..
have "IH.map = H"
using H.extensionality IH.extensionality H.functor_axioms by force
moreover have "IH.cmp = Φ⇩H"
proof
fix μν
show "IH.cmp μν = Φ⇩H μν"
using IH.cmp_def D.VV.arr_char⇩S⇩b⇩C D.VV.dom_simp D.VV.cod_simp H.FF_def
C.comp_arr_dom C.comp_cod_arr H.Φ.naturality1 H.Φ.extensionality
by (cases "D.VV.arr μν") auto
qed
ultimately show "pseudonatural_equivalence_whisker_left V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
H Φ⇩H GFH.map GFH.cmp K Φ⇩K ρoH.map⇩0 ρoH.map⇩1"
using KoρoH.pseudonatural_equivalence_whisker_left_axioms by simp
qed
interpretation ξoG: pseudonatural_equivalence_whisker_right V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
HK.FG.map HK.FG.cmp I⇩C.map I⇩C.cmp G Φ⇩G ξ⇩0 ξ⇩1
..
interpretation FoξoG: pseudonatural_equivalence_whisker_left V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
HKG.map HKG.cmp G Φ⇩G F Φ⇩F ξoG.map⇩0 ξoG.map⇩1
proof -
interpret FoξoG: pseudonatural_equivalence_whisker_left V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
HKG.map HKG.cmp IG.map IG.cmp F Φ⇩F ξoG.map⇩0 ξoG.map⇩1
..
have "IG.map = G"
using G_.extensionality IG.extensionality
by (meson G_.functor_axioms comp_identity_functor)
moreover have "IG.cmp = Φ⇩G"
proof
fix μν
show "IG.cmp μν = Φ⇩G μν"
using IG.cmp_def B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp G_.FF_def
C.comp_arr_dom C.comp_cod_arr G_.Φ.naturality1 G_.Φ.extensionality
by (cases "B.VV.arr μν") auto
qed
ultimately show "pseudonatural_equivalence_whisker_left V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
HKG.map HKG.cmp G Φ⇩G F Φ⇩F ξoG.map⇩0 ξoG.map⇩1"
using FoξoG.pseudonatural_equivalence_whisker_left_axioms by simp
qed
sublocale unit: composite_pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
I⇩D.map I⇩D.cmp HK.GF.map HK.GF.cmp KGFH.map KGFH.cmp
ζ⇩0 ζ⇩1 KoρoH.map⇩0 KoρoH.map⇩1
..
sublocale counit: composite_pseudonatural_equivalence
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
FHKG.map FHKG.cmp FG.FG.map FG.FG.cmp I⇩B.map I⇩B.cmp
FoξoG.map⇩0 FoξoG.map⇩1 σ⇩0 σ⇩1
..
abbreviation left_map
where "left_map ≡ FH.map"
abbreviation right_map
where "right_map ≡ KG.map"
abbreviation left_cmp
where "left_cmp ≡ FH.cmp"
abbreviation right_cmp
where "right_cmp ≡ KG.cmp"
definition unit⇩0
where "unit⇩0 ≡ unit.map⇩0"
definition unit⇩1
where "unit⇩1 ≡ unit.map⇩1"
definition counit⇩0
where "counit⇩0 ≡ counit.map⇩0"
definition counit⇩1
where "counit⇩1 == counit.map⇩1"
lemma unit⇩0_simp:
assumes "D.obj a"
shows "unit⇩0 a = K (ρ⇩0 (H.map⇩0 a)) ⋆⇩D ζ⇩0 a"
using assms unit⇩0_def unit.map⇩0_def KoρoH.map⇩0_def ρoH.map⇩0_def by simp
lemma unit⇩1_simp:
assumes "D.ide f"
shows "unit⇩1 f = 𝖺⇩D⇧-⇧1[K (ρ⇩0 (H.map⇩0 (trg⇩D f))), ζ⇩0 (trg⇩D f), f] ⋅⇩D
(K (ρ⇩0 (H.map⇩0 (trg⇩D f))) ⋆⇩D ζ⇩1 f) ⋅⇩D
𝖺⇩D[K (ρ⇩0 (H.map⇩0 (trg⇩D f))), K (H f), ζ⇩0 (src⇩D f)] ⋅⇩D
(D.inv (Φ⇩K (ρ⇩0 (H.map⇩0 (trg⇩D f)), H f)) ⋅⇩D
K (ρ⇩1 (H f)) ⋅⇩D
Φ⇩K (G (F (H f)), ρ⇩0 (H.map⇩0 (src⇩D f))) ⋆⇩D ζ⇩0 (src⇩D f)) ⋅⇩D
𝖺⇩D⇧-⇧1[K (G (F (H f))), K (ρ⇩0 (H.map⇩0 (src⇩D f))), ζ⇩0 (src⇩D f)]"
using assms unit⇩1_def unit.map⇩1_def KoρoH.map⇩0_def ρoH.map⇩0_def KoρoH.map⇩1_def ρoH.map⇩1_def
by simp
lemma counit⇩0_simp:
assumes "B.obj a"
shows "counit⇩0 a = σ⇩0 a ⋆⇩B F (ξ⇩0 (G_.map⇩0 a))"
using assms counit⇩0_def counit.map⇩0_def FoξoG.map⇩0_def ξoG.map⇩0_def by simp
lemma counit⇩1_simp:
assumes "B.ide f"
shows "counit⇩1 f = 𝖺⇩B⇧-⇧1[σ⇩0 (trg⇩B f), F (ξ⇩0 (G_.map⇩0 (trg⇩B f))), F (H (K (G f)))] ⋅⇩B
(σ⇩0 (trg⇩B f) ⋆⇩B
B.inv (Φ⇩F (ξ⇩0 (G_.map⇩0 (trg⇩B f)), H (K (G f)))) ⋅⇩B
F (ξ⇩1 (G f)) ⋅⇩B
Φ⇩F (G f, ξ⇩0 (G_.map⇩0 (src⇩B f)))) ⋅⇩B
𝖺⇩B[σ⇩0 (trg⇩B f), F (G f), F (ξ⇩0 (G_.map⇩0 (src⇩B f)))] ⋅⇩B
(σ⇩1 f ⋆⇩B F (ξ⇩0 (G_.map⇩0 (src⇩B f)))) ⋅⇩B
𝖺⇩B⇧-⇧1[f, σ⇩0 (src⇩B f), F (ξ⇩0 (G_.map⇩0 (src⇩B f)))]"
using assms counit⇩1_def counit.map⇩1_def FoξoG.map⇩0_def FoξoG.map⇩1_def
ξoG.map⇩0_def ξoG.map⇩1_def
by simp
sublocale equivalence_of_bicategories V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
FH.map FH.cmp KG.map KG.cmp unit⇩0 unit⇩1 counit⇩0 counit⇩1
proof -
interpret FH_KG: composite_pseudofunctor
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
KG.map KG.cmp FH.map FH.cmp
..
interpret KG_FH: composite_pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
FH.map FH.cmp KG.map KG.cmp
..
have "FH_KG.map = FHKG.map" by auto
moreover have "FH_KG.cmp = FHKG.cmp"
proof
fix μν
show "FH_KG.cmp μν = FHKG.cmp μν"
proof (cases "B.VV.arr μν")
case False
thus ?thesis
using FH_KG.Φ.extensionality FHKG.Φ.extensionality by simp
next
case True
have "FH_KG.cmp μν =
F (H (K (G (I⇩B.cmp μν)))) ⋅⇩B
F (H (K (G (B.dom (fst μν) ⋆⇩B B.dom (snd μν))))) ⋅⇩B
F (H (K (Φ⇩G (B.dom (fst μν), B.dom (snd μν))))) ⋅⇩B
F (H (Φ⇩K (G (B.dom (fst μν)), G (B.dom (snd μν))))) ⋅⇩B
(F (H (K (G (B.dom (fst μν))) ⋆⇩D K (G (B.dom (snd μν))))) ⋅⇩B
F (Φ⇩H (K (G (B.dom (fst μν))), K (G (B.dom (snd μν)))))) ⋅⇩B
Φ⇩F (H (K (G (B.dom (fst μν)))), H (K (G (B.dom (snd μν)))))"
using True FH_KG.cmp_def FHKG.cmp_def HKG.cmp_def
FH.cmp_def HK.FG.cmp_def KG.cmp_def
B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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⇩B.cmp μν)))) ⋅⇩B
F (H (K (G (B.dom (fst μν) ⋆⇩B B.dom (snd μν))))) ⋅⇩B
F (H (K (Φ⇩G (B.dom (fst μν), B.dom (snd μν))))) ⋅⇩B
F (H (Φ⇩K (G (B.dom (fst μν)), G (B.dom (snd μν))))) ⋅⇩B
F (Φ⇩H (K (G (B.dom (fst μν))), K (G (B.dom (snd μν))))) ⋅⇩B
Φ⇩F (H (K (G (B.dom (fst μν)))), H (K (G (B.dom (snd μν)))))"
proof -
have "F (H (K (G (B.dom (fst μν))) ⋆⇩D K (G (B.dom (snd μν))))) ⋅⇩B
F (Φ⇩H (K (G (B.dom (fst μν))), K (G (B.dom (snd μν))))) =
F (Φ⇩H (K (G (B.dom (fst μν))), K (G (B.dom (snd μν)))))"
using True B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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⇩B.cmp μν)))) ⋅⇩B
F (H (K (G (B.dom (fst μν) ⋆⇩B B.dom (snd μν))))) ⋅⇩B
F (H (K (Φ⇩G (B.dom (fst μν), B.dom (snd μν))))) ⋅⇩B
(F (H (K (G (B.dom (fst μν)) ⋆⇩C G (B.dom (snd μν))))) ⋅⇩B
F (H (Φ⇩K (G (B.dom (fst μν)), G (B.dom (snd μν)))))) ⋅⇩B
F (Φ⇩H (K (G (B.dom (fst μν))), K (G (B.dom (snd μν))))) ⋅⇩B
Φ⇩F (H (K (G (B.dom (fst μν)))), H (K (G (B.dom (snd μν)))))"
proof -
have "F (H (K (G (B.dom (fst μν)) ⋆⇩C G (B.dom (snd μν))))) ⋅⇩B
F (H (Φ⇩K (G (B.dom (fst μν)), G (B.dom (snd μν))))) =
F (H (Φ⇩K (G (B.dom (fst μν)), G (B.dom (snd μν)))))"
using True B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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 μν"
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⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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 μν
show "KG_FH.cmp μν = KGFH.cmp μν"
proof (cases "D.VV.arr μν")
case False
thus ?thesis
using KG_FH.Φ.extensionality KGFH.Φ.extensionality by simp
next
case True
have "KG_FH.cmp μν =
K (G (F (H (I⇩D.cmp μν)))) ⋅⇩D
K (G (F (H (D.dom (fst μν) ⋆⇩D D.dom (snd μν))))) ⋅⇩D
K (G (F (Φ⇩H (D.dom (fst μν), D.dom (snd μν))))) ⋅⇩D
K (G (Φ⇩F (H (D.dom (fst μν)), H (D.dom (snd μν))))) ⋅⇩D
(K (G (F (H (D.dom (fst μν))) ⋆⇩B F (H (D.dom (snd μν))))) ⋅⇩D
K (Φ⇩G (F (H (D.dom (fst μν))), F (H (D.dom (snd μν)))))) ⋅⇩D
Φ⇩K (G (F (H (D.dom (fst μν)))), G (F (H (D.dom (snd μν)))))"
using True KG_FH.cmp_def FH.cmp_def KG.cmp_def
B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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⇩D.cmp μν)))) ⋅⇩D
K (G (F (H (D.dom (fst μν) ⋆⇩D D.dom (snd μν))))) ⋅⇩D
K (G (F (Φ⇩H (D.dom (fst μν), D.dom (snd μν))))) ⋅⇩D
K (G (Φ⇩F (H (D.dom (fst μν)), H (D.dom (snd μν))))) ⋅⇩D
K (Φ⇩G (F (H (D.dom (fst μν))), F (H (D.dom (snd μν))))) ⋅⇩D
Φ⇩K (G (F (H (D.dom (fst μν)))), G (F (H (D.dom (snd μν)))))"
proof -
have "K (G (F (H (D.dom (fst μν))) ⋆⇩B F (H (D.dom (snd μν))))) ⋅⇩D
K (Φ⇩G (F (H (D.dom (fst μν))), F (H (D.dom (snd μν))))) =
K (Φ⇩G (F (H (D.dom (fst μν))), F (H (D.dom (snd μν)))))"
using True KG_FH.cmp_def FH.cmp_def KG.cmp_def
B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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⇩D.cmp μν)))) ⋅⇩D
K (G (F (H (D.dom (fst μν) ⋆⇩D D.dom (snd μν))))) ⋅⇩D
K (G (F (Φ⇩H (D.dom (fst μν), D.dom (snd μν))))) ⋅⇩D
(K (G (F (H (D.dom (fst μν)) ⋆⇩C H (D.dom (snd μν))))) ⋅⇩D
K (G (Φ⇩F (H (D.dom (fst μν)), H (D.dom (snd μν)))))) ⋅⇩D
K (Φ⇩G (F (H (D.dom (fst μν))), F (H (D.dom (snd μν))))) ⋅⇩D
Φ⇩K (G (F (H (D.dom (fst μν)))), G (F (H (D.dom (snd μν)))))"
proof -
have "K (G (F (H (D.dom (fst μν)) ⋆⇩C H (D.dom (snd μν))))) ⋅⇩D
K (G (Φ⇩F (H (D.dom (fst μν)), H (D.dom (snd μν))))) =
K (G (Φ⇩F (H (D.dom (fst μν)), H (D.dom (snd μν)))))"
using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def
B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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 μν"
using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def
B.VV.arr_char⇩S⇩b⇩C B.VV.dom_simp B.VV.cod_simp
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_simp C.VV.cod_simp
D.VV.arr_char⇩S⇩b⇩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⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
FH.map FH.cmp KG.map KG.cmp unit⇩0 unit⇩1 counit⇩0 counit⇩1"
unfolding unit⇩0_def unit⇩1_def counit⇩0_def counit⇩1_def
using B.bicategory_axioms D.bicategory_axioms FH.pseudofunctor_axioms
KG.pseudofunctor_axioms unit.pseudonatural_equivalence_axioms
counit.pseudonatural_equivalence_axioms I⇩B.identity_pseudofunctor_axioms
I⇩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⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
left_map left_cmp right_map right_cmp unit⇩0 unit⇩1 counit⇩0 counit⇩1"
..
end
subsection "Equivalence with a Dense Sub-bicategory"
text ‹
The purpose of this section is to show that, given a bicategory ‹B› and a sub-bicategory
defined by a ``dense'' set of objects of ‹B›, the embedding of the sub-bicategory in ‹B›
extends to an equivalence of bicategories. Here by ``dense'' we mean that every object
of ‹B› is equivalent to some object of the subbicategory.
›
locale dense_subbicategory =
B: bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B +
subbicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ‹λμ. B.arr μ ∧ src⇩B μ ∈ Obj ∧ trg⇩B μ ∈ Obj›
for V⇩B :: "'b comp" (infixr ‹⋅⇩B› 55)
and H⇩B :: "'b comp" (infixr ‹⋆⇩B› 53)
and 𝖺⇩B :: "'b ⇒ 'b ⇒ 'b ⇒ 'b" (‹𝖺⇩B[_, _, _]›)
and 𝗂⇩B :: "'b ⇒ 'b" (‹𝗂⇩B[_]›)
and src⇩B :: "'b ⇒ 'b"
and trg⇩B :: "'b ⇒ 'b"
and Obj :: "'b set" +
assumes dense: "⋀a. B.obj a ⟹ ∃a'. a' ∈ Obj ∧ B.equivalent_objects a' a"
begin
notation B.𝖺' (‹𝖺⇩B⇧-⇧1[_, _, _]›)
notation B.lunit (‹𝗅⇩B[_]›)
notation B.lunit' (‹𝗅⇩B⇧-⇧1[_]›)
notation B.runit (‹𝗋⇩B[_]›)
notation B.runit' (‹𝗋⇩B⇧-⇧1[_]›)
notation comp (infixr ‹⋅› 55)
notation hcomp (infixr ‹⋆› 53)
notation in_hom (‹«_ : _ ⇒ _»›)
notation in_hhom (‹«_ : _ → _»›)
notation 𝖺 (‹𝖺[_, _, _]›)
notation 𝖺' (‹𝖺⇧-⇧1[_, _, _]›)
notation lunit (‹𝗅[_]›)
notation lunit' (‹𝗅⇧-⇧1[_]›)
notation runit (‹𝗋[_]›)
notation runit' (‹𝗋⇧-⇧1[_]›)
abbreviation (input) Arr
where "Arr ≡ λμ. B.arr μ ∧ src⇩B μ ∈ Obj ∧ trg⇩B μ ∈ Obj"
sublocale emb: embedding_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B Arr
..
abbreviation E
where "E ≡ emb.map"
abbreviation Φ⇩E
where "Φ⇩E ≡ emb.cmp"
text ‹
We define a projection ‹P› by transporting arrows of ‹B› across chosen
equivalences between objects of ‹B› and objects of the sub-bicategory.
›
definition P⇩0
where "P⇩0 a ≡ SOME a'. obj a' ∧ B.equivalent_objects a' a"
lemma P⇩0_props:
assumes "B.obj a"
shows "obj (P⇩0 a)"
and "B.equivalent_objects (P⇩0 a) a"
and "B.equivalent_objects a a' ⟹ P⇩0 a = P⇩0 a'"
and "P⇩0 (P⇩0 a) = P⇩0 a"
and "B.obj (P⇩0 a)"
and "P⇩0 a ∈ Obj"
proof -
have "∃a'. obj a' ∧ B.equivalent_objects a' a"
using assms dense [of a] obj_char arr_char⇩S⇩b⇩C
by (metis (no_types, lifting) B.equivalent_objects_def B.in_hhomE B.obj_def B.obj_simps(3))
hence 1: "obj (P⇩0 a) ∧ B.equivalent_objects (P⇩0 a) a"
unfolding P⇩0_def
using someI_ex [of "λa'. obj a' ∧ B.equivalent_objects a' a"] by blast
show "obj (P⇩0 a)"
using 1 by simp
show 2: "B.equivalent_objects (P⇩0 a) a"
using 1 by simp
show 3: "⋀a'. B.equivalent_objects a a' ⟹ P⇩0 a = P⇩0 a'"
unfolding P⇩0_def
using B.equivalent_objects_symmetric B.equivalent_objects_transitive by meson
show "P⇩0 (P⇩0 a) = P⇩0 a"
using 2 3 [of "P⇩0 a"] B.equivalent_objects_symmetric by auto
show "B.obj (P⇩0 a)"
using 1 B.equivalent_objects_def by auto
thus "P⇩0 a ∈ Obj"
using 1 3 obj_char arr_char⇩S⇩b⇩C by auto
qed
text ‹
For each object ‹a› of ‹B›, we choose an adjoint equivalence from ‹a› to ‹P⇩0 a›.
The use of adjoint equivalences is necessary in order to establish the required
coherence conditions.
›
definition e
where "e a = (SOME e. «e : a →⇩B P⇩0 a» ∧
(∃d η ε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B e d η ε))"
definition d
where "d a = (SOME d. ∃η ε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) d η ε)"
definition η
where "η a = (SOME η. ∃ε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) η ε)"
definition ε
where "ε a = (SOME ε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) (η a) ε)"
lemma chosen_adjoint_equivalence:
assumes "B.obj a"
shows "adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) (η a) (ε a)"
and "«e a : a →⇩B P⇩0 a»" and "B.ide (d a)" and "B.ide (e a)" and "B.iso (η a)" and "B.iso (ε a)"
proof -
have "∃e. «e : a →⇩B P⇩0 a» ∧
(∃d η ε. adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B e d η ε)"
proof -
obtain e where e: "«e : a →⇩B P⇩0 a» ∧ B.equivalence_map e"
using assms P⇩0_props(2) B.equivalent_objects_symmetric B.equivalent_objects_def
by meson
obtain d η ε where d: "adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B e d η ε"
using e B.equivalence_map_extends_to_adjoint_equivalence by blast
thus ?thesis
using e d by auto
qed
hence 1: "«e a : a →⇩B P⇩0 a» ∧
adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) (η a) (ε a)"
using d_def e_def η_def ε_def arr_char⇩S⇩b⇩C
someI_ex [of "λe. «e : a →⇩B P⇩0 a» ∧
(∃d η ε. adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B e d η ε)"]
someI_ex [of "λd. (∃η ε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) d η ε)"]
someI_ex [of "λη. ∃ε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) η ε"]
someI_ex [of "λε. adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) (η a) ε"]
by simp
interpret adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ‹e a› ‹d a› ‹η a› ‹ε a›
using 1 by simp
show "adjoint_equivalence_in_bicategory
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B (e a) (d a) (η a) (ε a)" ..
show "«e a : a →⇩B P⇩0 a»" using 1 by simp
show "B.ide (d a)" by simp
show "B.ide (e a)" by simp
show "B.iso (η a)" by simp
show "B.iso (ε a)" by simp
qed
lemma equivalence_data_in_hom⇩B [intro]:
assumes "B.obj a"
shows "«e a : a →⇩B P⇩0 a»" and "«d a : P⇩0 a →⇩B a»"
and "«e a : e a ⇒⇩B e a»" and "«d a : d a ⇒⇩B d a»"
and "«η a : a →⇩B a»" and "«ε a : P⇩0 a →⇩B P⇩0 a»"
and "«η a : a ⇒⇩B d a ⋆⇩B e a»" and "«ε a : e a ⋆⇩B d a ⇒⇩B P⇩0 a»"
proof -
interpret adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ‹e a› ‹d a› ‹η a› ‹ε a›
using assms chosen_adjoint_equivalence by simp
show e: "«e a : a →⇩B P⇩0 a»"
using assms chosen_adjoint_equivalence by simp
show "«d a : P⇩0 a →⇩B a»" using e antipar by auto
show "«e a : e a ⇒⇩B e a»" by auto
show "«d a : d a ⇒⇩B d a»" by auto
show "«η a : a →⇩B a»" using unit_in_hom e by auto
show "«ε a : P⇩0 a →⇩B P⇩0 a»" using counit_in_hom e by auto
show "«η a : a ⇒⇩B d a ⋆⇩B e a»" using unit_in_hom e by auto
show "«ε a : e a ⋆⇩B d a ⇒⇩B P⇩0 a»" using counit_in_hom e by auto
qed
lemma equivalence_data_simps⇩B [simp]:
assumes "B.obj a"
shows "B.ide (d a)" and "B.ide (e a)" and "B.iso (η a)" and "B.iso (ε a)"
and "src⇩B (e a) = a" and "trg⇩B (e a) = P⇩0 a" and "src⇩B (d a) = P⇩0 a" and "trg⇩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⇩B (η a) = a" and "trg⇩B (η a) = a" and "src⇩B (ε a) = P⇩0 a" and "trg⇩B (ε a) = P⇩0 a"
and "B.dom (η a) = a" and "B.cod (η a) = d a ⋆⇩B e a"
and "B.dom (ε a) = e a ⋆⇩B d a" and "B.cod (ε a) = P⇩0 a"
using assms chosen_adjoint_equivalence equivalence_data_in_hom⇩B B.in_hhom_def
apply auto
by (meson B.in_homE)+
lemma equivalence_data_in_hom [intro]:
assumes "obj a"
shows "«e a : a → P⇩0 a»" and "«d a : P⇩0 a → a»"
and "«e a : e a ⇒ e a»" and "«d a : d a ⇒ d a»"
and "«η a : a → a»" and "«ε a : P⇩0 a → P⇩0 a»"
and "«η a : a ⇒ d a ⋆ e a»" and "«ε a : e a ⋆ d a ⇒ P⇩0 a»"
proof -
have a: "B.obj a ∧ a ∈ Obj"
using assms obj_char arr_char⇩S⇩b⇩C by auto
have P⇩0a: "obj (P⇩0 a) ∧ P⇩0 a ∈ Obj"
using assms a P⇩0_props [of a] arr_char⇩S⇩b⇩C obj_char by auto
show ea: "«e a : a → P⇩0 a»"
using a P⇩0a arr_char⇩S⇩b⇩C chosen_adjoint_equivalence(2) src_def trg_def by auto
show "«e a : e a ⇒ e a»"
using a ea chosen_adjoint_equivalence(4) ide_char⇩S⇩b⇩C in_hom_char⇩S⇩b⇩C by auto
show da: "«d a : P⇩0 a → a»"
using a P⇩0a arr_char⇩S⇩b⇩C src_def trg_def P⇩0_props(1) equivalence_data_in_hom⇩B(2)
by auto
show "«d a : d a ⇒ d a»"
using a da chosen_adjoint_equivalence(3) ide_char⇩S⇩b⇩C in_hom_char⇩S⇩b⇩C by auto
show ηa: "«η a : a ⇒ d a ⋆ e a»"
proof
show 1: "arr (η a)"
using assms a P⇩0a da ea arr_char⇩S⇩b⇩C equivalence_data_in_hom⇩B(7) hcomp_closed
by (metis (no_types, lifting) equivalence_data_in_hom⇩B(5) B.in_hhom_def)
show "dom (η a) = a"
using a 1 dom_char⇩S⇩b⇩C equivalence_data_in_hom⇩B(7) by auto
show "cod (η a) = d a ⋆ e a"
using a 1 da ea cod_char⇩S⇩b⇩C in_hom_char⇩S⇩b⇩C equivalence_data_in_hom⇩B(7) hcomp_def
by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE)
qed
show εa: "«ε a : e a ⋆ d a ⇒ P⇩0 a»"
proof
show 1: "arr (ε a)"
using assms a P⇩0a da ea arr_char⇩S⇩b⇩C equivalence_data_in_hom⇩B(8) hcomp_closed
by (metis (no_types, lifting) equivalence_data_in_hom⇩B(6) B.in_hhom_def)
show "dom (ε a) = e a ⋆ d a"
using a 1 da ea dom_char⇩S⇩b⇩C in_hom_char⇩S⇩b⇩C equivalence_data_in_hom⇩B(8) hcomp_def
by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE)
show "cod (ε a) = P⇩0 a"
using a 1 cod_char⇩S⇩b⇩C equivalence_data_in_hom⇩B(8) by auto
qed
show "«η a : a → a»"
using assms ηa src_def trg_def src_dom trg_dom
by (metis (no_types, lifting) in_hhom_def in_hom_char⇩S⇩b⇩C obj_simps(2-3)
vconn_implies_hpar(1-2))
show "«ε a : P⇩0 a → P⇩0 a»"
using P⇩0a εa src_def trg_def src_cod trg_cod
by (metis (no_types, lifting) in_hhom_def in_hom_char⇩S⇩b⇩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 (η a)" and "iso (ε a)"
and "src (e a) = a" and "trg (e a) = P⇩0 a" and "src (d a) = P⇩0 a" and "trg (d a) = a"
and "dom (e a) = e a" and "cod (e a) = e a"
and "dom (d a) = d a" and "cod (d a) = d a"
and "src (η a) = a" and "trg (η a) = a" and "src (ε a) = P⇩0 a" and "trg (ε a) = P⇩0 a"
and "dom (η a) = a" and "cod (η a) = d a ⋆⇩B e a"
and "dom (ε a) = e a ⋆⇩B d a" and "cod (ε a) = P⇩0 a"
using assms equivalence_data_in_hom(2) ide_char⇩S⇩b⇩C obj_char apply auto[1]
using assms equivalence_data_in_hom(1) ide_char⇩S⇩b⇩C obj_char apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C iso_char⇩S⇩b⇩C B.iso_is_arr apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C iso_char⇩S⇩b⇩C P⇩0_props B.iso_is_arr apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C P⇩0_props src_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C P⇩0_props trg_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C P⇩0_props src_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C P⇩0_props trg_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C P⇩0_props apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C P⇩0_props apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C P⇩0_props apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C P⇩0_props apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C B.iso_is_arr src_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C B.iso_is_arr trg_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C P⇩0_props B.iso_is_arr src_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C P⇩0_props B.iso_is_arr trg_def apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C B.iso_is_arr apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C B.iso_is_arr apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C P⇩0_props B.iso_is_arr apply auto[1]
using assms obj_char arr_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C P⇩0_props B.iso_is_arr by auto[1]
definition P
where "P μ = e (trg⇩B μ) ⋆⇩B μ ⋆⇩B d (src⇩B μ)"
lemma P_in_hom⇩B [intro]:
assumes "B.arr μ"
shows "«P μ : P⇩0 (src⇩B μ) →⇩B P⇩0 (trg⇩B μ)»"
and "«P μ : P (B.dom μ) ⇒⇩B P (B.cod μ)»"
unfolding P_def using assms by auto
lemma P_simps⇩B [simp]:
assumes "B.arr μ"
shows "B.arr (P μ)"
and "src⇩B (P μ) = P⇩0 (src⇩B μ)" and "trg⇩B (P μ) = P⇩0 (trg⇩B μ)"
and "B.dom (P μ) = P (B.dom μ)" and "B.cod (P μ) = P (B.cod μ)"
using assms P_in_hom⇩B by blast+
lemma P_in_hom [intro]:
assumes "B.arr μ"
shows "«P μ : P⇩0 (src⇩B μ) → P⇩0 (trg⇩B μ)»"
and "«P μ : P (B.dom μ) ⇒ P (B.cod μ)»"
proof -
show 1: "«P μ : P⇩0 (src⇩B μ) → P⇩0 (trg⇩B μ)»"
using assms P_in_hom⇩B(1) arr_char⇩S⇩b⇩C src_def trg_def P⇩0_props(1)
by (metis (no_types, lifting) in_hhom_def obj_char B.in_hhomE B.obj_simps(2)
B.obj_src B.obj_trg)
show "«P μ : P (B.dom μ) ⇒ P (B.cod μ)»"
using assms 1 dom_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C
by (intro in_homI) auto
qed
lemma P_simps [simp]:
assumes "B.arr μ"
shows "arr (P μ)"
and "src (P μ) = P⇩0 (src⇩B μ)" and "trg (P μ) = P⇩0 (trg⇩B μ)"
and "dom (P μ) = P (B.dom μ)" and "cod (P μ) = P (B.cod μ)"
using assms P_in_hom by blast+
interpretation P: "functor" V⇩B comp P
proof
show "⋀μ. ¬ B.arr μ ⟹ P μ = null"
using P_def null_char B.hseq_char B.hseq_char' by auto
have 0: "⋀μ. B.arr μ ⟹ B.arr (P μ)"
by simp
show 1: "⋀μ. B.arr μ ⟹ arr (P μ)"
using P_simps⇩B(1) by simp
show 2: "⋀μ. B.arr μ ⟹ dom (P μ) = P (B.dom μ)"
using 1 dom_simp by auto
show 3: "⋀μ. B.arr μ ⟹ cod (P μ) = P (B.cod μ)"
using 1 cod_simp by auto
show "⋀μ ν. B.seq μ ν ⟹ P (μ ⋅⇩B ν) = P μ ⋅ P ν"
proof -
fix μ ν
assume seq: "B.seq μ ν"
show "P (μ ⋅⇩B ν) = P μ ⋅ P ν"
proof -
have 4: "P (μ ⋅⇩B ν) = e (trg⇩B μ) ⋆⇩B μ ⋅⇩B ν ⋆⇩B d (src⇩B μ)"
unfolding P_def using seq B.src_vcomp B.trg_vcomp by simp
also have "... = P μ ⋅⇩B P ν"
unfolding P_def
using seq 0 4 B.whisker_left B.whisker_right
by (metis equivalence_data_simps⇩B(1-2) B.hseqE B.obj_trg B.seqE B.src_vcomp
B.vseq_implies_hpar(2))
also have "... = P μ ⋅ P ν"
using seq 1 seq_char⇩S⇩b⇩C comp_char by auto
finally show ?thesis by blast
qed
qed
qed
interpretation faithful_functor V⇩B comp P
proof
fix μ μ'
assume par: "B.par μ μ'"
have 1: "src⇩B μ = src⇩B μ' ∧ trg⇩B μ = trg⇩B μ'"
using par by (metis B.src_dom B.trg_dom)
assume eq: "P μ = P μ'"
interpret src: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (src⇩B μ)› ‹d (src⇩B μ)› ‹η (src⇩B μ)› ‹ε (src⇩B μ)›
using par chosen_adjoint_equivalence by simp
interpret trg: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (trg⇩B μ)› ‹d (trg⇩B μ)› ‹η (trg⇩B μ)› ‹ε (trg⇩B μ)›
using par chosen_adjoint_equivalence by simp
show "μ = μ'"
using eq par 1
unfolding P_def
using B.equivalence_cancel_left [of "e (trg⇩B μ)" "μ ⋆⇩B d (src⇩B μ)" "μ' ⋆⇩B d (src⇩B μ')"]
B.equivalence_cancel_right [of "d (src⇩B μ)" μ μ']
B.equivalence_map_def src.dual_equivalence trg.equivalence_in_bicategory_axioms
by auto
qed
interpretation P: weak_arrow_of_homs V⇩B src⇩B trg⇩B comp src trg P
proof
fix μ
assume μ: "B.arr μ"
interpret src: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (src⇩B μ)› ‹d (src⇩B μ)› ‹η (src⇩B μ)› ‹ε (src⇩B μ)›
using μ chosen_adjoint_equivalence by simp
show "isomorphic (P (src⇩B μ)) (src (P μ))"
proof (unfold isomorphic_def)
show "∃f. «f : P (src⇩B μ) ⇒ src (P μ)» ∧ iso f"
proof -
let ?φ = "ε (src⇩B μ) ⋅⇩B (e (src⇩B μ) ⋆⇩B 𝗅⇩B[d (src⇩B μ)])"
have "«?φ : P (src⇩B μ) ⇒ src (P μ)» ∧ iso ?φ"
proof -
have 1: "«?φ : P (src⇩B μ) ⇒⇩B P⇩0 (src⇩B μ)» ∧ B.iso ?φ"
unfolding P_def using μ by auto
moreover have "arr ?φ ∧ arr (B.inv ?φ)"
using μ 1 arr_char⇩S⇩b⇩C P⇩0_props(1)
by (metis (no_types, lifting) P⇩0_props(6) B.arrI B.arr_inv equivalence_data_simps⇩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⇩0 (src⇩B μ) = src (P μ)"
using μ by simp
ultimately show ?thesis
using μ in_hom_char⇩S⇩b⇩C arr_char⇩S⇩b⇩C iso_char⇩S⇩b⇩C
by (metis (no_types, lifting) src_closed equivalence_data_simps⇩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⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (trg⇩B μ)› ‹d (trg⇩B μ)› ‹η (trg⇩B μ)› ‹ε (trg⇩B μ)›
using μ chosen_adjoint_equivalence by simp
show "isomorphic (P (trg⇩B μ)) (trg (P μ))"
proof (unfold isomorphic_def)
show "∃f. «f : P (trg⇩B μ) ⇒ trg (P μ)» ∧ iso f"
proof -
let ?ψ = "ε (trg⇩B μ) ⋅⇩B (e (trg⇩B μ) ⋆⇩B 𝗅⇩B[d (trg⇩B μ)])"
have "«?ψ : P (trg⇩B μ) ⇒ trg (P μ)» ∧ iso ?ψ"
proof -
have 1: "«?ψ : P (trg⇩B μ) ⇒⇩B P⇩0 (trg⇩B μ)» ∧ B.iso ?ψ"
unfolding P_def using μ by auto
moreover have "arr ?ψ ∧ arr (B.inv ?ψ)"
using μ 1 arr_char⇩S⇩b⇩C P⇩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⇩B μ)) ∧ arr (P⇩0 (trg⇩B μ))"
using μ arr_char⇩S⇩b⇩C P⇩0_props(1) P_simps(1) obj_char by auto
moreover have "P⇩0 (trg⇩B μ) = trg (P μ)"
using μ by simp
ultimately show ?thesis
using in_hom_char⇩S⇩b⇩C iso_char⇩S⇩b⇩C by force
qed
thus ?thesis by blast
qed
qed
qed
text ‹
The following seems to be needed to avoid non-confluent simplifications,
\emph{e.g.}~of ‹S.src (P μ)› to ‹P.map⇩0 a› and to ‹P⇩0 a›.
›
lemma P_map⇩0_simp [simp]:
assumes "B.obj a"
shows "P.map⇩0 a = P⇩0 a"
using assms P.map⇩0_def B.obj_simps(1-2) by simp
interpretation HoPP: composite_functor B.VV.comp VV.comp comp
P.FF ‹λμν. hcomp (fst μν) (snd μν)›
..
interpretation PoH: composite_functor B.VV.comp V⇩B comp ‹(λμν. fst μν ⋆⇩B snd μν)› P
..
no_notation B.in_hom (‹«_ : _ →⇩B _»›)
definition CMP
where "CMP f g ≡ (e (trg⇩B f) ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d (src⇩B g)]) ⋅⇩B
(e (trg⇩B f) ⋆⇩B f ⋆⇩B
𝗅⇩B[g ⋆⇩B d (src⇩B g)] ⋅⇩B (B.inv (η (trg⇩B g)) ⋆⇩B g ⋆⇩B d (src⇩B g))) ⋅⇩B
(e (trg⇩B f) ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d (src⇩B f), e (trg⇩B g), g ⋆⇩B d (src⇩B g)]) ⋅⇩B
𝖺⇩B[e (trg⇩B f), f, d (src⇩B f) ⋆⇩B P g] ⋅⇩B
𝖺⇩B[e (trg⇩B f) ⋆⇩B f, d (src⇩B f), P g] ⋅⇩B
(𝖺⇩B⇧-⇧1[e (trg⇩B f), f, d (src⇩B f)] ⋆⇩B P g)"
text ‹
The 2-cell ‹CMP f g› has the right type to be a compositor for a pseudofunctor
whose underlying mapping is ‹P›.
›
lemma CMP_in_hom [intro]:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "«CMP f g : P⇩0 (src⇩B g) → P⇩0 (trg⇩B f)»"
and "«CMP f g : P f ⋆ P g ⇒ P (f ⋆⇩B g)»"
and "«CMP f g : P⇩0 (src⇩B g) →⇩B P⇩0 (trg⇩B f)»"
and "«CMP f g : P f ⋆⇩B P g ⇒⇩B P (f ⋆⇩B g)»"
proof -
show 1: "«CMP f g : P f ⋆⇩B P g ⇒⇩B P (f ⋆⇩B g)»"
using assms
by (unfold P_def CMP_def, intro B.comp_in_homI' B.seqI B.hseqI') auto
show 2: "«CMP f g : P⇩0 (src⇩B g) →⇩B P⇩0 (trg⇩B f)»"
using assms 1 B.src_cod B.trg_cod B.vconn_implies_hpar(1-2) by auto
show 3: "«CMP f g : P f ⋆ P g ⇒ P (f ⋆⇩B g)»"
using assms 1 B.arrI B.vconn_implies_hpar(1-2) P_simps(1) arr_char⇩S⇩b⇩C hcomp_eqI
hseqI' in_hom_char⇩S⇩b⇩C
by force
show "«CMP f g : P⇩0 (src⇩B g) → P⇩0 (trg⇩B f)»"
using 2 3 arr_char⇩S⇩b⇩C src_def trg_def by fastforce
qed
lemma CMP_simps [simp]:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "arr (CMP f g)"
and "src (CMP f g) = P⇩0 (src⇩B g)" and "trg (CMP f g) = P⇩0 (trg⇩B f)"
and "dom (CMP f g) = P f ⋆ P g" and "cod (CMP f g) = P (f ⋆⇩B g)"
and "B.arr (CMP f g)"
and "src⇩B (CMP f g) = P⇩0 (src⇩B g)" and "trg⇩B (CMP f g) = P⇩0 (trg⇩B f)"
and "B.dom (CMP f g) = P f ⋆⇩B P g" and "B.cod (CMP f g) = P (f ⋆⇩B g)"
using assms CMP_in_hom [of f g] by auto
lemma iso_CMP:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩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⇩S⇩b⇩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⇩S⇩b⇩C arr_char⇩S⇩b⇩C CMP_simps(1) by auto
qed
abbreviation (input) SRC
where "SRC μ ≡ d (src⇩B μ) ⋆⇩B e (src⇩B μ)"
abbreviation (input) TRG
where "TRG μ ≡ d (trg⇩B μ) ⋆⇩B e (trg⇩B μ)"
definition LUNIT
where "LUNIT f ≡ 𝗅⇩B[f] ⋅⇩B (B.inv (η (trg⇩B f)) ⋆⇩B f)"
definition RUNIT
where "RUNIT f ≡ 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η (src⇩B f)))"
text ‹
Here we prove a series of results that would be automatic if we had some notion of
``bicategory with ‹SRC› and ‹TRG› as alternative source and target''.
Perhaps this idea can be developed in future work and used to simplify the overall
development.
›
lemma LUNIT_in_hom [intro]:
assumes "B.ide f"
shows "«LUNIT f : src⇩B f →⇩B trg⇩B f»"
and "«LUNIT f : TRG f ⋆⇩B f ⇒⇩B f»"
proof -
interpret e_trg_f: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (trg⇩B f)› ‹d (trg⇩B f)› ‹η (trg⇩B f)› ‹ε (trg⇩B f)›
using assms chosen_adjoint_equivalence by simp
show "«LUNIT f : src⇩B f →⇩B trg⇩B f»"
unfolding LUNIT_def
using assms e_trg_f.unit_is_iso by auto
show "«LUNIT f : TRG f ⋆⇩B f ⇒⇩B f»"
unfolding LUNIT_def
using assms e_trg_f.unit_is_iso by auto
qed
lemma LUNIT_simps [simp]:
assumes "B.ide f"
shows "B.arr (LUNIT f)"
and "src⇩B (LUNIT f) = src⇩B f" and "trg⇩B (LUNIT f) = trg⇩B f"
and "B.dom (LUNIT f) = TRG f ⋆⇩B f"
and "B.cod (LUNIT f) = f"
using assms LUNIT_in_hom by auto
lemma RUNIT_in_hom [intro]:
assumes "B.ide f"
shows "«RUNIT f : src⇩B f →⇩B trg⇩B f»"
and "«RUNIT f : f ⋆⇩B SRC f ⇒⇩B f»"
proof -
interpret e_src_f: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (src⇩B f)› ‹d (src⇩B f)› ‹η (src⇩B f)› ‹ε (src⇩B f)›
using assms chosen_adjoint_equivalence by simp
show "«RUNIT f : src⇩B f →⇩B trg⇩B f»"
unfolding RUNIT_def
using assms e_src_f.unit_is_iso by auto
show "«RUNIT f : f ⋆⇩B SRC f ⇒⇩B f»"
unfolding RUNIT_def
using assms e_src_f.unit_is_iso by auto
qed
lemma RUNIT_simps [simp]:
assumes "B.ide f"
shows "B.arr (RUNIT f)"
and "src⇩B (RUNIT f) = src⇩B f" and "trg⇩B (RUNIT f) = trg⇩B f"
and "B.dom (RUNIT f) = f ⋆⇩B SRC f"
and "B.cod (RUNIT f) = f"
using assms RUNIT_in_hom by auto
lemma iso_LUNIT:
assumes "B.ide f"
shows "B.iso (LUNIT f)"
by (simp add: assms B.isos_compose LUNIT_def)
lemma iso_RUNIT:
assumes "B.ide f"
shows "B.iso (RUNIT f)"
by (simp add: assms B.isos_compose RUNIT_def)
lemma LUNIT_naturality:
assumes "B.arr μ"
shows "μ ⋅⇩B LUNIT (B.dom μ) = LUNIT (B.cod μ) ⋅⇩B (TRG μ ⋆⇩B μ)"
proof -
interpret e_trg_μ: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (trg⇩B μ)› ‹d (trg⇩B μ)› ‹η (trg⇩B μ)› ‹ε (trg⇩B μ)›
using assms chosen_adjoint_equivalence by simp
show ?thesis
proof -
have "μ ⋅⇩B LUNIT (B.dom μ) = (μ ⋅⇩B 𝗅⇩B[B.dom μ]) ⋅⇩B (B.inv (η (trg⇩B μ)) ⋆⇩B B.dom μ)"
unfolding LUNIT_def
using assms B.comp_assoc by simp
also have "... = 𝗅⇩B[B.cod μ] ⋅⇩B (trg⇩B μ ⋆⇩B μ) ⋅⇩B (B.inv (η (trg⇩B μ)) ⋆⇩B B.dom μ)"
using assms B.lunit_naturality B.comp_assoc by simp
also have "... = 𝗅⇩B[B.cod μ] ⋅⇩B (B.inv (η (trg⇩B μ)) ⋆⇩B μ)"
using assms B.interchange [of "trg⇩B μ" "B.inv (η (trg⇩B μ))" μ "B.dom μ"]
e_trg_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
also have "... = 𝗅⇩B[B.cod μ] ⋅⇩B (B.inv (η (trg⇩B μ)) ⋆⇩B B.cod μ) ⋅⇩B (TRG μ ⋆⇩B μ)"
using assms B.interchange [of "B.inv (η (trg⇩B μ))" "TRG μ" "B.cod μ" μ]
e_trg_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
also have "... = LUNIT (B.cod μ) ⋅⇩B (TRG μ ⋆⇩B μ)"
unfolding LUNIT_def
using assms B.comp_assoc by simp
finally show ?thesis by simp
qed
qed
lemma RUNIT_naturality:
assumes "B.arr μ"
shows "μ ⋅⇩B RUNIT (B.dom μ) = RUNIT (B.cod μ) ⋅⇩B (μ ⋆⇩B SRC μ)"
proof -
interpret e_src_μ: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (src⇩B μ)› ‹d (src⇩B μ)› ‹η (src⇩B μ)› ‹ε (src⇩B μ)›
using assms chosen_adjoint_equivalence by simp
show ?thesis
proof -
have "μ ⋅⇩B RUNIT (B.dom μ) =
(μ ⋅⇩B 𝗋⇩B[B.dom μ]) ⋅⇩B (B.dom μ ⋆⇩B B.inv (η (src⇩B μ)))"
unfolding RUNIT_def
using assms B.comp_assoc by simp
also have "... = 𝗋⇩B[B.cod μ] ⋅⇩B (μ ⋆⇩B src⇩B μ) ⋅⇩B (B.dom μ ⋆⇩B B.inv (η (src⇩B μ)))"
using assms B.runit_naturality B.comp_assoc by simp
also have "... = 𝗋⇩B[B.cod μ] ⋅⇩B (μ ⋆⇩B B.inv (η (src⇩B μ)))"
using assms B.interchange [of μ "B.dom μ" "src⇩B μ" "B.inv (η (src⇩B μ))"]
e_src_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
also have "... = 𝗋⇩B[B.cod μ] ⋅⇩B (B.cod μ ⋆⇩B B.inv (η (src⇩B μ))) ⋅⇩B (μ ⋆⇩B SRC μ)"
using assms B.interchange [of "B.cod μ" μ "B.inv (η (src⇩B μ))" "SRC μ"]
e_src_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr
by simp
also have "... = RUNIT (B.cod μ) ⋅⇩B (μ ⋆⇩B SRC μ)"
unfolding RUNIT_def
using assms B.comp_assoc by simp
finally show ?thesis by simp
qed
qed
lemma LUNIT_hcomp:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "LUNIT (f ⋆⇩B g) ⋅⇩B 𝖺⇩B[d (trg⇩B f) ⋆⇩B e (trg⇩B f), f, g] = LUNIT f ⋆⇩B g"
proof -
interpret e_trg_f: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (trg⇩B f)› ‹d (trg⇩B f)› ‹η (trg⇩B f)› ‹ε (trg⇩B f)›
using assms chosen_adjoint_equivalence by simp
have "LUNIT (f ⋆⇩B g) ⋅⇩B 𝖺⇩B[TRG f, f, g] =
𝗅⇩B[f ⋆⇩B g] ⋅⇩B (B.inv (η (trg⇩B f)) ⋆⇩B f ⋆⇩B g) ⋅⇩B 𝖺⇩B[TRG f, f, g]"
unfolding LUNIT_def
using assms B.comp_assoc by simp
also have "... = (𝗅⇩B[f ⋆⇩B g] ⋅⇩B 𝖺⇩B[trg⇩B f, f, g]) ⋅⇩B ((B.inv (η (trg⇩B f)) ⋆⇩B f) ⋆⇩B g)"
using assms B.assoc_naturality [of "B.inv (η (trg⇩B f))" f g] e_trg_f.unit_is_iso
B.comp_assoc
by simp
also have "... = (𝗅⇩B[f] ⋆⇩B g) ⋅⇩B ((B.inv (η (trg⇩B f)) ⋆⇩B f) ⋆⇩B g)"
using assms B.lunit_hcomp by simp
also have "... = LUNIT f ⋆⇩B g"
using assms LUNIT_def LUNIT_simps(1) B.whisker_right by auto
finally show ?thesis by simp
qed
lemma RUNIT_hcomp:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "RUNIT (f ⋆⇩B g) = (f ⋆⇩B RUNIT g) ⋅⇩B 𝖺⇩B[f, g, SRC g]"
proof -
interpret e_src_g: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (src⇩B g)› ‹d (src⇩B g)› ‹η (src⇩B g)› ‹ε (src⇩B g)›
using assms chosen_adjoint_equivalence by simp
have "(f ⋆⇩B RUNIT g) ⋅⇩B 𝖺⇩B[f, g, SRC g] =
(f ⋆⇩B 𝗋⇩B[g]) ⋅⇩B (f ⋆⇩B g ⋆⇩B B.inv (η (src⇩B g))) ⋅⇩B 𝖺⇩B[f, g, SRC g]"
unfolding RUNIT_def
using assms B.whisker_left e_src_g.unit_is_iso B.comp_assoc by simp
also have "... = ((f ⋆⇩B 𝗋⇩B[g]) ⋅⇩B 𝖺⇩B[f, g, src⇩B g]) ⋅⇩B ((f ⋆⇩B g) ⋆⇩B B.inv (η (src⇩B g)))"
using assms B.assoc_naturality [of f g "B.inv (η (src⇩B g))"] e_src_g.unit_is_iso
B.comp_assoc
by simp
also have "... = 𝗋⇩B[f ⋆⇩B g] ⋅⇩B ((f ⋆⇩B g) ⋆⇩B B.inv (η (src⇩B g)))"
using assms B.runit_hcomp by simp
also have "... = RUNIT (f ⋆⇩B g)"
using assms RUNIT_def by simp
finally show ?thesis by simp
qed
lemma TRIANGLE:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "(f ⋆⇩B LUNIT g) ⋅⇩B 𝖺⇩B[f, SRC f, g] = RUNIT f ⋆⇩B g"
proof -
interpret e_trg_g: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e (trg⇩B g)› ‹d (trg⇩B g)› ‹η (trg⇩B g)› ‹ε (trg⇩B g)›
using assms chosen_adjoint_equivalence by simp
show ?thesis
proof -
have "(f ⋆⇩B LUNIT g) ⋅⇩B 𝖺⇩B[f, SRC f, g] =
(f ⋆⇩B 𝗅⇩B[g]) ⋅⇩B (f ⋆⇩B B.inv (η (trg⇩B g)) ⋆⇩B g) ⋅⇩B 𝖺⇩B[f, SRC f, g]"
using assms B.whisker_left e_trg_g.unit_is_iso LUNIT_def LUNIT_simps(1) B.comp_assoc
by auto
also have "... = ((f ⋆⇩B 𝗅⇩B[g]) ⋅⇩B 𝖺⇩B[f, src⇩B f, g]) ⋅⇩B ((f ⋆⇩B B.inv (η (trg⇩B g))) ⋆⇩B g)"
using assms B.assoc_naturality [of f "B.inv (η (trg⇩B g))" g] e_trg_g.unit_is_iso
B.comp_assoc
by auto
also have "... = (𝗋⇩B[f] ⋆⇩B g) ⋅⇩B ((f ⋆⇩B B.inv (η (trg⇩B g))) ⋆⇩B g)"
using assms B.triangle by simp
also have "... = RUNIT f ⋆⇩B g"
using assms B.whisker_right e_trg_g.unit_is_iso RUNIT_def RUNIT_simps
by metis
finally show ?thesis by simp
qed
qed
text ‹
The ‹CMP f g› also satisfy the naturality conditions required of compositors.
›
lemma CMP_naturality:
assumes "B.arr μ" and "B.arr ν" and "src⇩B μ = trg⇩B ν"
shows "CMP (B.cod μ) (B.cod ν) ⋅⇩B (P μ ⋆⇩B P ν)
= P (μ ⋆⇩B ν) ⋅⇩B CMP (B.dom μ) (B.dom ν)"
proof -
let ?a = "src⇩B ν"
let ?b = "src⇩B μ"
let ?c = "trg⇩B μ"
let ?f = "B.dom μ"
let ?g = "B.cod μ"
let ?h = "B.dom ν"
let ?k = "B.cod ν"
have "CMP ?g ?k ⋅⇩B (P μ ⋆⇩B P ν)
= (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, ?g, d ?b ⋆⇩B P ?k] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B ?g, d ?b, P ?k] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?g, d ?b] ⋆⇩B P ?k) ⋅⇩B
(P μ ⋆⇩B P ν)"
unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
also have "... = (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, ?g, d ?b ⋆⇩B P ?k] ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B ?g, d ?b, P ?k] ⋅⇩B
(((e ?c ⋆⇩B μ) ⋆⇩B d ?b) ⋆⇩B P ν)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
proof -
have "(𝖺⇩B⇧-⇧1[e ?c, ?g, d ?b] ⋆⇩B P ?k) ⋅⇩B (P μ ⋆⇩B P ν)
= 𝖺⇩B⇧-⇧1[e ?c, ?g, d ?b] ⋅⇩B P μ ⋆⇩B P ?k ⋅⇩B P ν"
using assms P_def B.interchange by fastforce
also have
"... = ((e ?c ⋆⇩B μ) ⋆⇩B d ?b) ⋅⇩B 𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?k ⋅⇩B P ν"
using assms P_def B.assoc'_naturality [of "e ?c" μ "d ?b"] by simp
also have
"... = ((e ?c ⋆⇩B μ) ⋆⇩B d ?b) ⋅⇩B 𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ν ⋅⇩B P ?h"
using assms B.comp_arr_dom B.comp_cod_arr B.src_cod B.src_dom
B.trg_cod B.trg_dom P.as_nat_trans.naturality
by simp
also have "... = (((e ?c ⋆⇩B μ) ⋆⇩B d ?b) ⋆⇩B P ν) ⋅⇩B (𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
using assms B.interchange by auto
finally have "(𝖺⇩B⇧-⇧1[e ?c, ?g, d ?b] ⋆⇩B P ?k) ⋅⇩B (P μ ⋆⇩B P ν)
= (((e ?c ⋆⇩B μ) ⋆⇩B d ?b) ⋆⇩B P ν) ⋅⇩B (𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B
(𝖺⇩B[e ?c, ?g, d ?b ⋆⇩B P ?k] ⋅⇩B
((e ?c ⋆⇩B μ) ⋆⇩B d ?b ⋆⇩B P ν)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B ?f, d ?b, P ?h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
using assms B.assoc_naturality [of "e ?c ⋆⇩B μ" "d ?b" "P ν"] B.comp_assoc by simp
also have "... = (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B
(e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B
((e ?c ⋆⇩B ?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B
(e ?c ⋆⇩B μ ⋆⇩B d ?b ⋆⇩B P ν)) ⋅⇩B
𝖺⇩B[e ?c, ?f, d ?b ⋆⇩B P ?h] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B ?f, d ?b, P ?h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
using assms B.assoc_naturality [of "e ?c" μ "d ?b ⋆⇩B P ν"] B.comp_assoc by simp
also have "... = (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B
((e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B μ ⋆⇩B SRC μ ⋆⇩B ν ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, ?f, d ?b ⋆⇩B P ?h] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B ?f, d ?b, P ?h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
proof -
have
"(e ?c ⋆⇩B ?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B (e ?c ⋆⇩B μ ⋆⇩B d ?b ⋆⇩B P ν)
= (e ?c ⋆⇩B μ ⋆⇩B TRG ν ⋆⇩B ν ⋆⇩B d ?a) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a])"
proof -
have "(e ?c ⋆⇩B ?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B (e ?c ⋆⇩B μ ⋆⇩B d ?b ⋆⇩B P ν)
= e ?c ⋆⇩B (?g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]) ⋅⇩B (μ ⋆⇩B d ?b ⋆⇩B P ν)"
using assms P_def B.whisker_left by simp
also have "... = e ?c ⋆⇩B μ ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a] ⋅⇩B (d ?b ⋆⇩B P ν)"
using assms P_def B.comp_cod_arr
B.interchange [of "?g" μ "𝖺⇩B⇧-⇧1[d ?b, e ?b, ?k ⋆⇩B d ?a]" "d ?b ⋆⇩B P ν"]
by simp
also have "... = e ?c ⋆⇩B μ ⋆⇩B
(TRG ν ⋆⇩B ν ⋆⇩B d ?a) ⋅⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a]"
using assms P_def B.assoc'_naturality [of "d ?b" "e ?b" "ν ⋆⇩B d ?a"]
by simp
also have "... = e ?c ⋆⇩B
(μ ⋆⇩B TRG ν ⋆⇩B ν ⋆⇩B d ?a) ⋅⇩B (?f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a])"
using assms B.comp_arr_dom
B.interchange [of μ "?f" "TRG ν ⋆⇩B ν ⋆⇩B d ?a""𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a]"]
by fastforce
also have
"... = (e ?c ⋆⇩B μ ⋆⇩B TRG ν ⋆⇩B ν ⋆⇩B d ?a) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a])"
using assms B.whisker_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = ((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B
(e ?c ⋆⇩B μ ⋆⇩B ν ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B LUNIT (?h ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, ?f, d ?b ⋆⇩B P ?h] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B ?f, d ?b, P ?h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
proof -
have "((e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B (e ?c ⋆⇩B μ ⋆⇩B TRG ν ⋆⇩B ν ⋆⇩B d ?a))
= e ?c ⋆⇩B μ ⋆⇩B LUNIT (B.cod (ν ⋆⇩B d ?a)) ⋅⇩B ((d ?b ⋆⇩B e ?b) ⋆⇩B ν ⋆⇩B d ?a)"
using assms comp_arr_dom B.comp_cod_arr B.whisker_left
B.interchange [of "?g" μ "LUNIT (?k ⋆⇩B d ?a)" "(d ?b ⋆⇩B e ?b) ⋆⇩B ν ⋆⇩B d ?a"]
by fastforce
also have "... = e ?c ⋆⇩B μ ⋆⇩B (ν ⋆⇩B d ?a) ⋅⇩B LUNIT (?h ⋆⇩B d ?a)"
using assms LUNIT_naturality [of "ν ⋆⇩B d ?a"] by simp
also have "... = (e ?c ⋆⇩B μ ⋆⇩B ν ⋆⇩B d ?a) ⋅⇩B (e ?c ⋆⇩B ?f ⋆⇩B LUNIT (?h ⋆⇩B d ?a))"
using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left
B.interchange [of μ "?f" "ν ⋆⇩B d ?a" "LUNIT (?h ⋆⇩B d ?a)"]
by simp
finally have "((e ?c ⋆⇩B ?g ⋆⇩B LUNIT (?k ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B μ ⋆⇩B TRG ν ⋆⇩B ν ⋆⇩B d ?a))
= (e ?c ⋆⇩B μ ⋆⇩B ν ⋆⇩B d ?a) ⋅⇩B (e ?c ⋆⇩B ?f ⋆⇩B LUNIT (?h ⋆⇩B d ?a))"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = P (μ ⋆⇩B ν) ⋅⇩B
(e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?f, ?h, d ?a]) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B LUNIT (?h ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B ?f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ?h ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, ?f, d ?b ⋆⇩B P ?h] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B ?f, d ?b, P ?h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, ?f, d ?b] ⋆⇩B P ?h)"
proof -
have "(e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B (e ?c ⋆⇩B μ ⋆⇩B ν ⋆⇩B d ?a)
= P (μ ⋆⇩B ν) ⋅⇩B (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?f, ?h, d ?a])"
proof -
have "(e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a]) ⋅⇩B (e ?c ⋆⇩B μ ⋆⇩B ν ⋆⇩B d ?a)
= e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?g, ?k, d ?a] ⋅⇩B (μ ⋆⇩B ν ⋆⇩B d ?a)"
using assms B.whisker_left by simp
also have "... = e ?c ⋆⇩B ((μ ⋆⇩B ν) ⋆⇩B d ?a) ⋅⇩B 𝖺⇩B⇧-⇧1[?f, ?h, d ?a]"
using assms B.assoc'_naturality [of μ ν "d ?a"] by simp
also have "... = P (μ ⋆⇩B ν) ⋅⇩B (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[?f, ?h, d ?a])"
using assms P_def B.whisker_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P (μ ⋆⇩B ν) ⋅⇩B CMP ?f ?h"
unfolding CMP_def LUNIT_def using assms by simp
finally show ?thesis by simp
qed
interpretation EV: self_evaluation_map V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ..
notation EV.eval (‹⦃_⦄›)
abbreviation (input) SRCt (‹❙S❙R❙C›)
where "❙S❙R❙C μ ≡ ❙⟨d (src⇩B μ)❙⟩ ❙⋆ ❙⟨e (src⇩B μ)❙⟩"
abbreviation (input) TRGt (‹❙T❙R❙G›)
where "❙T❙R❙G μ ≡ ❙⟨d (trg⇩B μ)❙⟩ ❙⋆ ❙⟨e (trg⇩B μ)❙⟩"
abbreviation (input) PRJt (‹❙P❙R❙J›)
where "❙P❙R❙J μ ≡ ❙⟨e (trg⇩B μ)❙⟩ ❙⋆ ❙⟨μ❙⟩ ❙⋆ ❙⟨d (src⇩B μ)❙⟩"
text ‹
The ‹CMP f g› satisfy the coherence conditions with respect to associativity that are
required of compositors.
›
lemma CMP_coherence:
assumes "B.ide f" and "B.ide g" and "B.ide h" and "src⇩B f = trg⇩B g" and "src⇩B g = trg⇩B h"
shows "P 𝖺⇩B[f, g, h] ⋅⇩B CMP (f ⋆⇩B g) h ⋅⇩B (CMP f g ⋆⇩B P h)
= CMP f (g ⋆⇩B h) ⋅⇩B (P f ⋆⇩B CMP g h) ⋅⇩B 𝖺⇩B[P f, P g, P h]"
proof -
text ‹
The overall strategy of the proof is to expand the definition of ‹CMP› on the
left and right-hand sides, then permute the occurrences of ‹LUNIT› and
‹RUNIT› to the left ends of both the left-hand side and right-hand side of the
equation to be proved, so that the initial portions of these expressions become
identical and the remaining parts to the right consist only of canonical isomorphisms.
Then the Coherence Theorem is applied to prove syntactically (and automatically) that the
canonical parts are equal, which implies equality of the complete expressions.
The rest is just grinding through the calculations.
›
let ?a = "trg⇩B f"
let ?b = "trg⇩B g"
let ?c = "trg⇩B h"
let ?d = "src⇩B h"
have "P 𝖺⇩B[f, g, h] ⋅⇩B CMP (f ⋆⇩B g) h ⋅⇩B (CMP f g ⋆⇩B P h)
= P 𝖺⇩B[f, g, h] ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT (h ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g)
⋆⇩B P h)"
unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
also have "... = P 𝖺⇩B[f, g, h] ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT (h ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
using assms B.whisker_right P_def by simp
also have "... = P 𝖺⇩B[f, g, h] ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT h ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT (h ⋆⇩B d ?d)
= e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B (LUNIT h ⋆⇩B d ?d) ⋅⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]"
using assms LUNIT_hcomp [of h "d ?d"] B.invert_side_of_triangle by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT h ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
using assms B.whisker_left by simp
finally have "e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT (h ⋆⇩B d ?d)
= (e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT h ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (P 𝖺⇩B[f, g, h] ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, h, d ?d]) ⋅⇩B (e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT h ⋆⇩B d ?d)
= e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, h, d ?d] ⋅⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B (((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]"
using assms B.assoc'_naturality [of "f ⋆⇩B g" "LUNIT h" "d ?d"] by simp
also have "... = (e ?a ⋆⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B LUNIT h ⋆⇩B d ?d)
= (e ?a ⋆⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "P 𝖺⇩B[f, g, h] ⋅⇩B (e ?a ⋆⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d)
= e ?a ⋆⇩B 𝖺⇩B[f, g, h] ⋅⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d"
using assms B.whisker_left B.whisker_right P_def by simp
also have "... = e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋅⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d"
using assms B.assoc_naturality [of f g "LUNIT h"] by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d)"
using assms B.whisker_left B.whisker_right by simp
finally have "P 𝖺⇩B[f, g, h] ⋅⇩B (e ?a ⋆⇩B ((f ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d)"
by simp
thus ?thesis using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
(((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B d ?c) ⋆⇩B P h)) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "((e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋆⇩B P h)
= (e ?a ⋆⇩B (RUNIT f ⋆⇩B g ⋆⇩B d ?c) ⋅⇩B 𝖺⇩B⇧-⇧1[f, TRG g, g ⋆⇩B d ?c]) ⋆⇩B P h"
using assms TRIANGLE [of f "g ⋆⇩B d ?c"] B.invert_side_of_triangle by simp
also have "... = ((e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B d ?c) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g, g ⋆⇩B d ?c]) ⋆⇩B P h)"
using assms B.whisker_left B.whisker_right P_def by simp
finally have "((e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B d ?c)) ⋆⇩B P h)
= ((e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B d ?c) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g, g ⋆⇩B d ?c]) ⋆⇩B P h)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋆⇩B P h)) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B TRG g, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B ((e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B d ?c) ⋆⇩B P h)
= (e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c] ⋅⇩B (RUNIT f ⋆⇩B g ⋆⇩B d ?c)) ⋆⇩B P h"
using assms B.whisker_left B.whisker_right P_def by simp
also have "... = (e ?a ⋆⇩B
((RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋅⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B TRG g, g, d ?c])
⋆⇩B P h"
using assms B.assoc'_naturality [of "RUNIT f" g "d ?c"] by simp
also have "... = ((e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B TRG g, g, d ?c]) ⋆⇩B P h)"
using assms B.whisker_left B.whisker_right P_def by simp
finally have "((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B d ?c) ⋆⇩B P h)
= ((e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B TRG g, g, d ?c]) ⋆⇩B P h)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f ⋆⇩B g, d ?c, P h] ⋅⇩B
(((e ?a ⋆⇩B (RUNIT f ⋆⇩B g)) ⋆⇩B d ?c) ⋆⇩B P h)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B ((e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋆⇩B P h)
= 𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋅⇩B (e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋆⇩B P h"
using assms B.whisker_left B.whisker_right P_def by simp
also have "... = ((e ?a ⋆⇩B (RUNIT f ⋆⇩B g)) ⋆⇩B d ?c) ⋅⇩B
𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h"
using assms B.assoc'_naturality [of "e ?a" "RUNIT f ⋆⇩B g" "d ?c"] by simp
also have "... = (((e ?a ⋆⇩B (RUNIT f ⋆⇩B g)) ⋆⇩B d ?c) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h)"
using assms B.whisker_left B.whisker_right P_def by simp
finally have "(𝖺⇩B⇧-⇧1[e ?a, f ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c) ⋆⇩B P h)
= (((e ?a ⋆⇩B (RUNIT f ⋆⇩B g)) ⋆⇩B d ?c) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(𝖺⇩B[e ?a, f ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
((e ?a ⋆⇩B (RUNIT f ⋆⇩B g)) ⋆⇩B d ?c ⋆⇩B P h)) ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B d ?b ⋆⇩B e ?b, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, d ?b ⋆⇩B e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
using assms B.assoc_naturality [of "e ?a ⋆⇩B (RUNIT f ⋆⇩B g)" "d ?c" "P h"] B.comp_assoc
by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c ⋆⇩B P h)) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
using assms B.assoc_naturality [of "e ?a" "RUNIT f ⋆⇩B g" "d ?c ⋆⇩B P h"] B.comp_assoc
by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d])) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "((e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c ⋆⇩B P h))
= e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d] ⋅⇩B (d ?c ⋆⇩B P h)"
using assms B.comp_cod_arr B.whisker_left P_def
B.interchange [of "f ⋆⇩B g" "RUNIT f ⋆⇩B g"]
by simp
also have "... = e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]"
using assms B.comp_arr_dom P_def by simp
finally have "((e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B d ?c ⋆⇩B P h))
= e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d])) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B
𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) =
e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B
𝖺⇩B⇧-⇧1[TRG h, h, d ?d] ⋅⇩B
𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]"
using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left
B.interchange
[of "f ⋆⇩B g" "RUNIT f ⋆⇩B g" "𝖺⇩B⇧-⇧1[TRG h, h, d ?d]" "𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]"]
by simp
also have "... = (e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d])"
using assms B.comp_arr_dom B.whisker_left
B.interchange [of "RUNIT f ⋆⇩B g" "(f ⋆⇩B SRC f) ⋆⇩B g" "𝖺⇩B⇧-⇧1[TRG h, h, d ?d]"
"𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]"]
by simp
finally have "(e ?a ⋆⇩B (f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d])
= (e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B
𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B[f, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d])
= e ?a ⋆⇩B (𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d]"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B
𝖺⇩B⇧-⇧1[f, g ⋆⇩B (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d] ⋅⇩B
(f ⋆⇩B 𝖺⇩B⇧-⇧1[g, (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d]) ⋅⇩B
𝖺⇩B[f, g, ((d ?c ⋆⇩B e ?c) ⋆⇩B h) ⋆⇩B d ?d]"
proof -
have "(𝖺⇩B⇧-⇧1[f, g, (d ?c ⋆⇩B e ?c) ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[f, g ⋆⇩B (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d] ⋅⇩B
(f ⋆⇩B 𝖺⇩B⇧-⇧1[g, (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d])
= 𝖺⇩B⇧-⇧1[f ⋆⇩B g, (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d] ⋅⇩B
𝖺⇩B⇧-⇧1[f, g, ((d ?c ⋆⇩B e ?c) ⋆⇩B h) ⋆⇩B d ?d]"
using assms B.pentagon' B.comp_assoc by simp
moreover have "B.inv (𝖺⇩B⇧-⇧1[f, g, (d ?c ⋆⇩B e ?c) ⋆⇩B h] ⋆⇩B d ?d)
= 𝖺⇩B[f, g, (d ?c ⋆⇩B e ?c) ⋆⇩B h] ⋆⇩B d ?d"
using assms by simp
ultimately show ?thesis
using assms B.comp_assoc
B.invert_opposite_sides_of_square
[of "𝖺⇩B⇧-⇧1[f, g, (d ?c ⋆⇩B e ?c) ⋆⇩B h] ⋆⇩B d ?d"
"𝖺⇩B⇧-⇧1[f, g ⋆⇩B (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d] ⋅⇩B
(f ⋆⇩B 𝖺⇩B⇧-⇧1[g, (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d])"
"𝖺⇩B⇧-⇧1[f ⋆⇩B g, (d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d]"
"𝖺⇩B⇧-⇧1[f, g, ((d ?c ⋆⇩B e ?c) ⋆⇩B h) ⋆⇩B d ?d]"]
by simp
qed
also have "... = (e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B 𝖺⇩B[f, g, TRG h ⋆⇩B h] ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B g, TRG h ⋆⇩B h, d ?d])
= (e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, (d ?c ⋆⇩B e ?c) ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B 𝖺⇩B[f, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])
= e ?a ⋆⇩B 𝖺⇩B[f, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d] ⋅⇩B ((RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B
(RUNIT f ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]"
using assms B.hseqI' B.assoc_naturality [of "RUNIT f" g "𝖺⇩B⇧-⇧1[TRG h, h, d ?d]"]
by simp
also have "... = (e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B 𝖺⇩B[f, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (RUNIT f ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])
= (e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d))) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c ⋆⇩B e ?c, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c ⋆⇩B e ?c, h, d ?d]
= (e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
using assms B.whisker_left B.comp_arr_dom B.comp_cod_arr
B.interchange [of "RUNIT f" "f ⋆⇩B SRC f" "g ⋆⇩B ((TRG h ⋆⇩B h) ⋆⇩B d ?d)"
"g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]"]
by simp
also have "... = (e ?a ⋆⇩B
(f ⋆⇩B LUNIT (g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
using assms TRIANGLE [of f "g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d"] by simp
also have "... = (e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
using assms B.whisker_left B.comp_assoc by simp
finally have "e ?a ⋆⇩B RUNIT f ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]
= (e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B f ⋆⇩B
(LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[d ?b ⋆⇩B e ?b, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]"
using assms LUNIT_hcomp [of g "(TRG h ⋆⇩B h) ⋆⇩B d ?d"]
B.invert_side_of_triangle
by simp
also have "... = (e ?a ⋆⇩B f ⋆⇩B LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d])"
using assms B.whisker_left by simp
finally have "e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B f ⋆⇩B LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B (LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B f ⋆⇩B
𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d] ⋅⇩B (LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B f ⋆⇩B
((LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d]"
using assms B.assoc'_naturality [of "LUNIT g" "TRG h ⋆⇩B h" "d ?d"]
by simp
also have "... = (e ?a ⋆⇩B f ⋆⇩B (LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B f ⋆⇩B (LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = ((e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B SRC g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B (LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d] ⋅⇩B
(f ⋆⇩B (LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B
((f ⋆⇩B LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d]"
using assms B.assoc'_naturality [of f "LUNIT g ⋆⇩B TRG h ⋆⇩B h" "d ?d"] by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B (LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
proof -
have "(e ?a ⋆⇩B (f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B TRG h ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B LUNIT h) ⋆⇩B d ?d"
using assms B.whisker_left B.whisker_right B.comp_arr_dom B.comp_cod_arr
B.interchange [of g "LUNIT g" "LUNIT h" "(d ?c ⋆⇩B e ?c) ⋆⇩B h"]
by simp
thus ?thesis
using assms by simp
qed
finally have L: "P 𝖺⇩B[f, g, h] ⋅⇩B CMP (f ⋆⇩B g) h ⋅⇩B (CMP f g ⋆⇩B P h)
= (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
by simp
have "CMP f (g ⋆⇩B h) ⋅⇩B (P f ⋆⇩B CMP g h) ⋅⇩B 𝖺⇩B[P f, P g, P h]
= (e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT ((g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
(P f ⋆⇩B
(e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B
(e ?b ⋆⇩B g ⋆⇩B LUNIT (h ⋆⇩B d ?d)) ⋅⇩B
(e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h)) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
also have "... = (e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT ((g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B LUNIT (h ⋆⇩B d ?d)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
using assms B.whisker_left P_def B.comp_assoc by auto
also have "... = ((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b ⋆⇩B e ?b, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B LUNIT (h ⋆⇩B d ?d)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "e ?a ⋆⇩B f ⋆⇩B LUNIT ((g ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B f ⋆⇩B (LUNIT (g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]"
using assms LUNIT_hcomp [of "g ⋆⇩B h" "d ?d"] B.invert_side_of_triangle
by simp
also have "... = (e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "e ?a ⋆⇩B f ⋆⇩B LUNIT ((g ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B LUNIT (h ⋆⇩B d ?d)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B h, d ?d]) ⋅⇩B (e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B h, d ?d] ⋅⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h) ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B ((f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]"
using assms B.assoc'_naturality [of f "LUNIT (g ⋆⇩B h)" "d ?d"]
LUNIT_in_hom [of "g ⋆⇩B h"]
by auto
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B LUNIT (g ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
((P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B RUNIT g ⋆⇩B h ⋆⇩B d ?d)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B LUNIT (h ⋆⇩B d ?d)
= P f ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h ⋆⇩B d ?d) ⋅⇩B 𝖺⇩B⇧-⇧1[g, TRG h, h ⋆⇩B d ?d]"
using assms TRIANGLE [of g "h ⋆⇩B d ?d"] B.invert_side_of_triangle by simp
also have "... = (P f ⋆⇩B e ?b ⋆⇩B RUNIT g ⋆⇩B h ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h, h ⋆⇩B d ?d])"
using assms B.whisker_left P_def by simp
finally have "P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B LUNIT (h ⋆⇩B d ?d) =
(P f ⋆⇩B e ?b ⋆⇩B RUNIT g ⋆⇩B h ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, TRG h, h ⋆⇩B d ?d])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B (P f ⋆⇩B e ?b ⋆⇩B RUNIT g ⋆⇩B h ⋆⇩B d ?d)
= P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d] ⋅⇩B (RUNIT g ⋆⇩B h ⋆⇩B d ?d)"
using assms B.whisker_left P_def by simp
also have "... = P f ⋆⇩B e ?b ⋆⇩B
((RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]"
using assms B.assoc'_naturality [of "RUNIT g" h "d ?d"] by auto
also have "... = (P f ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d])"
using assms B.whisker_left P_def by simp
finally have "(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B RUNIT g ⋆⇩B h ⋆⇩B d ?d)
= (P f ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P (g ⋆⇩B h)] ⋅⇩B
(((e ?a ⋆⇩B f) ⋆⇩B d ?b) ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B (P f ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)
= 𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d"
using assms B.comp_arr_dom B.comp_cod_arr P_def
B.interchange
[of "𝖺⇩B⇧-⇧1[e ?a, f, d ?b]" "P f" "P (g ⋆⇩B h)" "e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d"]
by simp
also have "... = (((e ?a ⋆⇩B f) ⋆⇩B d ?b) ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d)"
using assms B.comp_arr_dom B.comp_cod_arr
B.interchange
[of "(e ?a ⋆⇩B f) ⋆⇩B d ?b" "𝖺⇩B⇧-⇧1[e ?a, f, d ?b]"
"e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d"
"e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d"]
by simp
finally have "(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P (g ⋆⇩B h)) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)
= (((e ?a ⋆⇩B f) ⋆⇩B d ?b) ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b ⋆⇩B e ?b, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P (g ⋆⇩B h)] ⋅⇩B
((e ?a ⋆⇩B f) ⋆⇩B d ?b ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
using assms P_def B.comp_assoc
B.assoc_naturality [of "e ?a ⋆⇩B f" "d ?b" "e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d"]
by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (d ?b ⋆⇩B e ?b) ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b ⋆⇩B e ?b, g ⋆⇩B h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B d ?b ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?c ⋆⇩B e ?c, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?c ⋆⇩B e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
using assms P_def B.comp_assoc
B.assoc_naturality [of "e ?a" f "d ?b ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d"]
by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B TRG g ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have
"(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B d ?b ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B f ⋆⇩B
𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d] ⋅⇩B
(d ?b ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B f ⋆⇩B
(TRG g ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]"
using assms B.assoc'_naturality [of "d ?b" "e ?b" "(RUNIT g ⋆⇩B h) ⋆⇩B d ?d"] by auto
also have "... = (e ?a ⋆⇩B f ⋆⇩B SRC f ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d])"
using assms B.whisker_left by simp
finally have
"(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, (g ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B d ?b ⋆⇩B e ?b ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B f ⋆⇩B SRC f ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B TRG g ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)
= e ?a ⋆⇩B f ⋆⇩B
𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d] ⋅⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B f ⋆⇩B
((TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]"
using assms B.assoc'_naturality [of "TRG g" "RUNIT g ⋆⇩B h" "d ?d"] by simp
also have "... = (e ?a ⋆⇩B f ⋆⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B TRG g ⋆⇩B (RUNIT g ⋆⇩B h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B f ⋆⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d])"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d)
= e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d] ⋅⇩B
(f ⋆⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d)"
using assms B.whisker_left by simp
also have "... = e ?a ⋆⇩B
((f ⋆⇩B SRC f ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]"
using assms
B.assoc'_naturality [of f "(d ?b ⋆⇩B e ?b) ⋆⇩B (RUNIT g ⋆⇩B h)" "d ?d"]
by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d])"
using assms B.whisker_left by simp
finally have "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B g ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B (TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d)
= (e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d])"
using assms by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B 𝖺⇩B[g, d ?c ⋆⇩B e ?c, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d
= e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B(g ⋆⇩B LUNIT h) ⋅⇩B 𝖺⇩B[g, SRC g, h]) ⋆⇩B d ?d"
using assms TRIANGLE [of g h] by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B 𝖺⇩B[g, TRG h, h]) ⋆⇩B d ?d)"
using assms B.whisker_left B.whisker_right by simp
finally have "e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B (RUNIT g ⋆⇩B h)) ⋆⇩B d ?d
= (e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B 𝖺⇩B[g, TRG h, h]) ⋆⇩B d ?d)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
((e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B 𝖺⇩B[g, SRC g, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d
= e ?a ⋆⇩B (f ⋆⇩B (LUNIT g ⋆⇩B h) ⋅⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h]) ⋆⇩B d ?d"
using assms LUNIT_hcomp [of g h] B.invert_side_of_triangle by simp
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h]) ⋆⇩B d ?d)"
using assms B.whisker_left B.whisker_right by simp
finally have "e ?a ⋆⇩B (f ⋆⇩B LUNIT (g ⋆⇩B h)) ⋆⇩B d ?d
= (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h]) ⋆⇩B d ?d)"
by simp
thus ?thesis
using assms B.comp_assoc by simp
qed
also have "... = ((e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B (TRG g ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d)) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, TRG h ⋆⇩B h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B 𝖺⇩B[g, SRC g, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have
"(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d)
= e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h] ⋅⇩B (TRG g ⋆⇩B g ⋆⇩B LUNIT h)) ⋆⇩B d ?d"
using assms B.whisker_left B.whisker_right by auto
also have "... = e ?a ⋆⇩B
(f ⋆⇩B ((TRG g ⋆⇩B g) ⋆⇩B LUNIT h) ⋅⇩B
𝖺⇩B⇧-⇧1[TRG g, g, SRC g ⋆⇩B h]) ⋆⇩B d ?d"
using assms B.assoc'_naturality [of "TRG g" g "LUNIT h"] by auto
also have "... = (e ?a ⋆⇩B (f ⋆⇩B (TRG g ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, TRG h ⋆⇩B h]) ⋆⇩B d ?d)"
using assms B.whisker_left B.whisker_right by auto
finally have "(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B g ⋆⇩B LUNIT h) ⋆⇩B d ?d)
= (e ?a ⋆⇩B (f ⋆⇩B (TRG g ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, TRG h ⋆⇩B h]) ⋆⇩B d ?d)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, TRG h ⋆⇩B h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B TRG g ⋆⇩B 𝖺⇩B[g, d ?c ⋆⇩B e ?c, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B d ?c ⋆⇩B e ?c) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
proof -
have "(e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B (TRG g ⋆⇩B g) ⋆⇩B LUNIT h) ⋆⇩B d ?d)
= e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B LUNIT h) ⋆⇩B d ?d"
using assms B.whisker_left B.whisker_right B.comp_arr_dom B.comp_cod_arr
B.interchange [of "LUNIT g" "TRG g ⋆⇩B g" h "LUNIT h"]
by simp
thus ?thesis
using assms by simp
qed
finally have R: "CMP f (g ⋆⇩B h) ⋅⇩B (P f ⋆⇩B CMP g h) ⋅⇩B 𝖺⇩B[P f, P g, P h]
= (e ?a ⋆⇩B (f ⋆⇩B LUNIT g ⋆⇩B LUNIT h) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, TRG h ⋆⇩B h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B 𝖺⇩B[g, TRG h, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
using assms by simp
text ‹
The portions of the expressions on the right-hand sides of assertions ‹L› and ‹R›
that are not identical only involve canonical isomorphisms, and thus they can be proved
equal automatically by the simplifier, once we have expressed them in the formal
language of ‹B›.
›
let ?LHS = "(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, (TRG g ⋆⇩B g) ⋆⇩B TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g ⋆⇩B g, TRG h ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f, d ?b ⋆⇩B e ?b, g ⋆⇩B (TRG h ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[TRG h, h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B[f ⋆⇩B SRC f, g, TRG h ⋆⇩B h ⋆⇩B d ?d]) ⋅⇩B
(e ?a ⋆⇩B ((f ⋆⇩B SRC f) ⋆⇩B g) ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c ⋆⇩B P h] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B (f ⋆⇩B SRC f) ⋆⇩B g, d ?c, P h] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, (f ⋆⇩B SRC f) ⋆⇩B g, d ?c] ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f ⋆⇩B SRC f, g, d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, SRC f, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
((e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, g ⋆⇩B d ?c]) ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a, f, d ?b ⋆⇩B P g] ⋆⇩B P h) ⋅⇩B
(𝖺⇩B[e ?a ⋆⇩B f, d ?b, P g] ⋆⇩B P h) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B P g) ⋆⇩B P h)"
let ?LHSt = "(❙⟨e ?a❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨f❙⟩, (❙T❙R❙G g ❙⋆ ❙⟨g❙⟩) ❙⋆ ❙T❙R❙G h ❙⋆ ❙⟨h❙⟩, ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙T❙R❙G g ❙⋆ ❙⟨g❙⟩, ❙T❙R❙G h ❙⋆ ❙⟨h❙⟩, ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙T❙R❙G g, ❙⟨g❙⟩, (❙T❙R❙G h ❙⋆ ❙⟨h❙⟩) ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙𝖺❙[❙⟨f❙⟩, ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩, ❙⟨g❙⟩ ❙⋆ (❙T❙R❙G h ❙⋆ ❙⟨h❙⟩) ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ (❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f) ❙⋆ ❙⟨g❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙T❙R❙G h, ❙⟨h❙⟩, ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙𝖺❙[❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f, ❙⟨g❙⟩, ❙T❙R❙G h ❙⋆ ❙⟨h❙⟩ ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ((❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f) ❙⋆ ❙⟨g❙⟩) ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨d ?c❙⟩, ❙⟨e ?c❙⟩, ❙⟨h❙⟩ ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
❙𝖺❙[❙⟨e ?a❙⟩, (❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f) ❙⋆ ❙⟨g❙⟩, ❙⟨d ?c❙⟩ ❙⋆ ❙P❙R❙J h❙] ❙⋅
❙𝖺❙[❙⟨e ?a❙⟩ ❙⋆ (❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f) ❙⋆ ❙⟨g❙⟩, ❙⟨d ?c❙⟩, ❙P❙R❙J h❙] ❙⋅
(❙𝖺⇧-⇧1❙[❙⟨e ?a❙⟩, (❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f) ❙⋆ ❙⟨g❙⟩, ❙⟨d ?c❙⟩❙] ❙⋆ ❙P❙R❙J h) ❙⋅
((❙⟨e ?a❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f, ❙⟨g❙⟩, ❙⟨d ?c❙⟩❙]) ❙⋆ ❙P❙R❙J h) ❙⋅
((❙⟨e ?a❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨f❙⟩, ❙S❙R❙C f, ❙⟨g❙⟩ ❙⋆ ❙⟨d ?c❙⟩❙]) ❙⋆ ❙P❙R❙J h) ❙⋅
((❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨d ?b❙⟩, ❙⟨e ?b❙⟩, ❙⟨g❙⟩ ❙⋆ ❙⟨d ?c❙⟩❙]) ❙⋆ ❙P❙R❙J h) ❙⋅
(❙𝖺❙[❙⟨e ?a❙⟩, ❙⟨f❙⟩, ❙⟨d ?b❙⟩ ❙⋆ ❙P❙R❙J g❙] ❙⋆ ❙P❙R❙J h) ❙⋅
(❙𝖺❙[❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩, ❙⟨d ?b❙⟩, ❙P❙R❙J g❙] ❙⋆ ❙P❙R❙J h) ❙⋅
((❙𝖺⇧-⇧1❙[❙⟨e ?a❙⟩, ❙⟨f❙⟩, ❙⟨d ?b❙⟩❙] ❙⋆ ❙P❙R❙J g) ❙⋆ ❙P❙R❙J h)"
let ?RHS = "(e ?a ⋆⇩B (f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, g, TRG h ⋆⇩B h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B (f ⋆⇩B SRC f ⋆⇩B 𝖺⇩B[g, TRG h, h]) ⋆⇩B d ?d) ⋅⇩B
(e ?a ⋆⇩B 𝖺⇩B⇧-⇧1[f, TRG g ⋆⇩B (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[TRG g, (g ⋆⇩B SRC g) ⋆⇩B h, d ?d]) ⋅⇩B
(e ?a ⋆⇩B f ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d]) ⋅⇩B
𝖺⇩B[e ?a, f, d ?b ⋆⇩B P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
𝖺⇩B[e ?a ⋆⇩B f, d ?b, P ((g ⋆⇩B SRC g) ⋆⇩B h)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?a, f, d ?b] ⋆⇩B e ?b ⋆⇩B ((g ⋆⇩B SRC g) ⋆⇩B h) ⋆⇩B d ?d) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B SRC g, h, d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B 𝖺⇩B⇧-⇧1[g, SRC g, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B e ?b ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?c, e ?c, h ⋆⇩B d ?d]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b, g, d ?c ⋆⇩B P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B g, d ?c, P h]) ⋅⇩B
(P f ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, g, d ?c] ⋆⇩B P h) ⋅⇩B
𝖺⇩B[P f, P g, P h]"
let ?RHSt = "(❙⟨e ?a❙⟩ ❙⋆ (❙⟨f❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙T❙R❙G g, ❙⟨g❙⟩, ❙T❙R❙G h ❙⋆ ❙⟨h❙⟩❙]) ❙⋆ ❙⟨d ?d❙⟩) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ (❙⟨f❙⟩ ❙⋆ ❙S❙R❙C f ❙⋆ ❙𝖺❙[❙⟨g❙⟩, ❙T❙R❙G h, ❙⟨h❙⟩❙]) ❙⋆ ❙⟨d ?d❙⟩) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨f❙⟩, ❙T❙R❙G g ❙⋆ (❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g) ❙⋆ ❙⟨h❙⟩, ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙T❙R❙G g, (❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g) ❙⋆ ❙⟨h❙⟩, ❙⟨d ?d❙⟩❙]) ❙⋅
(❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨d ?b❙⟩, ❙⟨e ?b❙⟩, ((❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g) ❙⋆ ❙⟨h❙⟩) ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
❙𝖺❙[❙⟨e ?a❙⟩, ❙⟨f❙⟩, ❙⟨d ?b❙⟩ ❙⋆ (❙⟨e ?b❙⟩ ❙⋆ ((❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g) ❙⋆ ❙⟨h❙⟩) ❙⋆ ❙⟨d ?d❙⟩)❙] ❙⋅
❙𝖺❙[❙⟨e ?a❙⟩ ❙⋆ ❙⟨f❙⟩, ❙⟨d ?b❙⟩, (❙⟨e ?b❙⟩ ❙⋆ ((❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g) ❙⋆ ❙⟨h❙⟩) ❙⋆ ❙⟨d ?d❙⟩)❙] ❙⋅
(❙𝖺⇧-⇧1❙[❙⟨e ?a❙⟩, ❙⟨f❙⟩, ❙⟨d ?b❙⟩❙] ❙⋆ ❙⟨e ?b❙⟩ ❙⋆ ((❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g) ❙⋆ ❙⟨h❙⟩) ❙⋆ ❙⟨d ?d❙⟩) ❙⋅
(❙P❙R❙J f ❙⋆ ❙⟨e ?b❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨g❙⟩ ❙⋆ ❙S❙R❙C g, ❙⟨h❙⟩, ❙⟨d ?d❙⟩❙]) ❙⋅
(❙P❙R❙J f ❙⋆ ❙⟨e ?b❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨g❙⟩, ❙S❙R❙C g, ❙⟨h❙⟩ ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
(❙P❙R❙J f ❙⋆ ❙⟨e ?b❙⟩ ❙⋆ ❙⟨g❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨d ?c❙⟩, ❙⟨e ?c❙⟩, ❙⟨h❙⟩ ❙⋆ ❙⟨d ?d❙⟩❙]) ❙⋅
(❙P❙R❙J f ❙⋆ ❙𝖺❙[❙⟨e ?b❙⟩, ❙⟨g❙⟩, ❙⟨d ?c❙⟩ ❙⋆ ❙P❙R❙J h❙]) ❙⋅
(❙P❙R❙J f ❙⋆ ❙𝖺❙[❙⟨e ?b❙⟩ ❙⋆ ❙⟨g❙⟩, ❙⟨d ?c❙⟩, ❙P❙R❙J h❙]) ❙⋅
(❙P❙R❙J f ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨e ?b❙⟩, ❙⟨g❙⟩, ❙⟨d ?c❙⟩❙] ❙⋆ ❙P❙R❙J h) ❙⋅
❙𝖺❙[❙P❙R❙J f, ❙P❙R❙J g, ❙P❙R❙J h❙]"
have E: "?LHS = ?RHS"
proof -
have "?LHS = ⦃?LHSt⦄"
using assms B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def
by simp
also have "... = ⦃?RHSt⦄"
using assms by (intro EV.eval_eqI, auto)
also have "... = ?RHS"
using assms B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩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
‹λμν. CMP (fst μν) (snd μν)›
proof
show 1: "⋀fg. B.VV.ide fg ⟹ «CMP (fst fg) (snd fg) : HoPP.map fg ⇒ PoH.map fg»"
using CMP_in_hom(2) B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P.FF_def hcomp_def arr_char⇩S⇩b⇩C
P⇩0_props(1) P.preserves_arr
by auto
show "⋀fg. B.VV.arr fg ⟹
CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) ⋅ HoPP.map fg =
PoH.map fg ⋅ CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
proof -
fix fg
assume fg: "B.VV.arr fg"
have "CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) ⋅ HoPP.map fg =
CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) ⋅⇩B HoPP.map fg"
using fg 1 B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C B.VV.cod_char⇩S⇩b⇩C HoPP.preserves_arr
P.FF_def hcomp_char comp_char
by auto
also have "... = PoH.map fg ⋅⇩B CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
using fg CMP_naturality [of "fst fg" "snd fg"]
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C B.VV.dom_char⇩S⇩b⇩C B.VV.cod_char⇩S⇩b⇩C P.FF_def
hcomp_def arr_char⇩S⇩b⇩C P⇩0_props(1) P.preserves_arr
by auto
also have "... = PoH.map fg ⋅ CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
proof -
have "P⇩0 (src⇩B (snd fg)) ∈ Obj ∧ P⇩0 (trg⇩B (fst fg)) ∈ Obj"
using fg P⇩0_props(6) B.VV.arrE B.obj_src B.obj_trg by meson
moreover have "B.seq (P (fst fg ⋆⇩B snd fg))
(CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg)))"
using fg 1 B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C B.VV.dom_char⇩S⇩b⇩C by simp
ultimately show ?thesis
using fg 1 comp_char arr_char⇩S⇩b⇩C in_hom_char⇩S⇩b⇩C by simp
qed
finally show "CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) ⋅ HoPP.map fg =
PoH.map fg ⋅ CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
by blast
qed
qed
interpretation CMP: natural_isomorphism B.VV.comp comp HoPP.map PoH.map CMP.map
using iso_CMP B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C CMP.map_simp_ide
by unfold_locales simp
definition Φ⇩P
where "Φ⇩P = CMP.map"
interpretation Φ⇩P: natural_isomorphism B.VV.comp comp HoPP.map PoH.map Φ⇩P
unfolding Φ⇩P_def
using CMP.natural_isomorphism_axioms by simp
no_notation B.in_hom (‹«_ : _ →⇩B _»›)
lemma Φ⇩P_in_hom [intro]:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "«Φ⇩P (f, g) : src (P g) → trg (P f)»"
and "«Φ⇩P (f, g) : P f ⋆ P g ⇒ P (f ⋆⇩B g)»"
proof -
show 1: "«Φ⇩P (f, g) : P f ⋆ P g ⇒ P (f ⋆⇩B g)»"
using assms B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P.FF_def by auto
show "«Φ⇩P (f, g) : src (P g) → trg (P f)»"
using 1
by (metis (no_types, lifting) hcomp_simps(2) in_hhom_def in_hom_char⇩S⇩b⇩C
src_hcomp vconn_implies_hpar(1-2))
qed
lemma Φ⇩P_simps [simp]:
assumes "B.ide f" and "B.ide g" and "src⇩B f = trg⇩B g"
shows "arr (Φ⇩P (f, g))"
and "src (Φ⇩P (f, g)) = src (P g)"
and "trg (Φ⇩P (f, g)) = trg (P f)"
and "dom (Φ⇩P (f, g)) = P f ⋆ P g"
and "cod (Φ⇩P (f, g)) = P (f ⋆⇩B g)"
using assms Φ⇩P_in_hom by blast+
sublocale prj: pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B comp hcomp 𝖺 𝗂⇩B src trg P Φ⇩P
proof
fix f g h
assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h"
and fg: "src⇩B f = trg⇩B g" and gh: "src⇩B g = trg⇩B h"
have 1: "ide (P f) ∧ ide (P g) ∧ ide (P h)"
using f g h P.preserves_ide by simp
have "P 𝖺⇩B[f, g, h] ⋅ Φ⇩P (f ⋆⇩B g, h) ⋅ (Φ⇩P (f, g) ⋆ P h) =
P 𝖺⇩B[f, g, h] ⋅⇩B Φ⇩P (f ⋆⇩B g, h) ⋅⇩B (Φ⇩P (f, g) ⋆⇩B P h)"
using f g h fg gh B.VV.arr_char⇩S⇩b⇩C B.VV.dom_char⇩S⇩b⇩C B.VV.cod_char⇩S⇩b⇩C P.FF_def
by (intro comp_eqI hcomp_eqI seqI hseqI') auto
also have "... = CMP f (g ⋆⇩B h) ⋅⇩B (P f ⋆⇩B CMP g h) ⋅⇩B 𝖺⇩B[P f, P g, P h]"
using f g h fg gh CMP_coherence CMP.map_simp_ide B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C
Φ⇩P_def
by simp
also have "... = Φ⇩P (f, g ⋆⇩B h) ⋅⇩B (P f ⋆⇩B Φ⇩P (g, h)) ⋅⇩B 𝖺⇩B[P f, P g, P h]"
using f g h fg gh B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C Φ⇩P_def by simp
also have "... = Φ⇩P (f, g ⋆⇩B h) ⋅ (P f ⋆⇩B Φ⇩P (g, h)) ⋅ 𝖺⇩B[P f, P g, P h]"
proof -
have 2: "arr ((P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h])"
using f g h fg gh B.VV.arr_char⇩S⇩b⇩C B.VV.dom_char⇩S⇩b⇩C P.FF_def by auto
moreover have "(P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h] =
(P f ⋆⇩B Φ⇩P (g, h)) ⋅⇩B 𝖺⇩B[P f, P g, P h]"
using f g h fg gh 2 comp_eqI hcomp_eqI assoc_simp by auto
moreover have 3: "arr (P f ⋆⇩B Φ⇩P (g, h))"
using f g h fg gh arr_hcompI
by (intro arr_hcompI hseqI') auto
moreover have "B.dom (P f ⋆⇩B Φ⇩P (g, h)) = P f ⋆⇩B P g ⋆⇩B P h"
proof -
have "B.dom (P f ⋆⇩B Φ⇩P (g, h)) = P f ⋆⇩B B.dom (Φ⇩P (g, h))"
using f g h fg gh 3 B.hcomp_simps(3)
by (metis (no_types, lifting) 1 arrE ideD(1) ideD(2) dom_char⇩S⇩b⇩C)
also have "... = P f ⋆⇩B P g ⋆⇩B P h"
using g h gh dom_char⇩S⇩b⇩C hcomp_char Φ⇩P_simps(1-4) by auto
finally show ?thesis by blast
qed
moreover have "B.dom (Φ⇩P (f, g ⋆⇩B h)) =
B.cod ((P f ⋆⇩B Φ⇩P (g, h)) ⋅⇩B 𝖺⇩B[P f, P g, P h])"
proof -
have "B.dom (Φ⇩P (f, g ⋆⇩B h)) = dom (Φ⇩P (f, g ⋆⇩B h))"
using f g h fg gh B.VV.arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C by simp
also have "... = cod ((P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h])"
using f g h fg gh 2 VV.arr_char⇩S⇩b⇩C Φ⇩P_simps by simp
also have "... = B.cod ((P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h])"
using 2 cod_char⇩S⇩b⇩C by presburger
also have "... = B.cod ((P f ⋆⇩B Φ⇩P (g, h)) ⋅⇩B 𝖺⇩B[P f, P g, P h])"
proof -
have "(P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h] = (P f ⋆⇩B Φ⇩P (g, h)) ⋅⇩B 𝖺⇩B[P f, P g, P h]"
using f g h fg gh 2 comp_eqI hcomp_eqI assoc_simp by auto
thus ?thesis by presburger
qed
finally show ?thesis by blast
qed
moreover have "B.cod 𝖺⇩B[P f, P g, P h] = P f ⋆⇩B P g ⋆⇩B P h"
using f g h fg gh 1 ide_char⇩S⇩b⇩C by auto
ultimately show ?thesis
using f g h fg gh B.VV.arr_char⇩S⇩b⇩C assoc_simp assoc_simps(1) dom_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C
by (intro comp_eqI' seqI arr_compI arr_hcompI) auto
qed
also have "... = Φ⇩P (f, g ⋆⇩B h) ⋅ (P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h]"
using f g h fg gh assoc_simp hcomp_eqI Φ⇩P_simps(1) by auto
finally show "P 𝖺⇩B[f, g, h] ⋅ Φ⇩P (f ⋆⇩B g, h) ⋅ (Φ⇩P (f, g) ⋆ P h) =
Φ⇩P (f, g ⋆⇩B h) ⋅ (P f ⋆ Φ⇩P (g, h)) ⋅ 𝖺[P f, P g, P h]"
by blast
qed
lemma pseudofunctor_prj:
shows "pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B comp hcomp 𝖺 𝗂⇩B src trg P Φ⇩P"
..
text ‹
We need an explicit formula for the unit constraints for ‹P›.
›
lemma prj_unit_char:
assumes "B.obj a"
shows "prj.unit a = (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)"
proof -
text ‹
We eventually need one of the triangle identities from the following interpretation.
However in the meantime, it contains a lot of simps that make an otherwise arduous
calculation much easier. So interpret it up front.
›
interpret adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ‹e a› ‹d a› ‹η a› ‹ε a›
using assms chosen_adjoint_equivalence(1) by simp
let ?x = "(e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)"
have x: "«?x : P.map⇩0 a ⇒⇩B P a»"
using assms P_def P.map⇩0_def P_map⇩0_simp
by (intro B.comp_in_homI') auto
have "?x = prj.unit a"
proof (intro prj.unit_eqI)
show "B.obj a" by fact
have a_da: "«a ⋆⇩B d a : P⇩0 a →⇩B a» ∧ B.ide (a ⋆⇩B d a)"
using assms B.obj_simps by auto
show x⇩S: "«?x : P.map⇩0 a ⇒ P a»"
using assms x P_map⇩0_simp P_simps⇩B(1) arr_char⇩S⇩b⇩C B.arrI
equivalence_data_simps⇩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⇩S⇩b⇩C)
show "iso ?x"
using assms x⇩S B.isos_compose iso_char⇩S⇩b⇩C arr_char⇩S⇩b⇩C by auto
have *: "«?x : P⇩0 a → P⇩0 a»"
using assms x⇩S P⇩0_props vconn_implies_hpar(1-2)
by (intro in_hhomI) auto
show "?x ⋅ 𝗂⇩B[P.map⇩0 a] = (P 𝗂⇩B[a] ⋅ Φ⇩P (a, a)) ⋅ (?x ⋆ ?x)"
proof -
have 0: "«d a ⋆⇩B e a ⋆⇩B a ⋆⇩B d a : P⇩0 a →⇩B a»"
using assms by force
have 1: "B.seq (B.inv (η a) ⋆⇩B a ⋆⇩B d a) 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"
using assms by (elim B.objE, intro B.seqI) auto
have "(P 𝗂⇩B[a] ⋅ Φ⇩P (a, a)) ⋅ (?x ⋆ ?x) = P 𝗂⇩B[a] ⋅ Φ⇩P (a, a) ⋅ (?x ⋆ ?x)"
using comp_assoc by simp
also have "... = P 𝗂⇩B[a] ⋅ (CMP a a ⋅ (P a ⋆ P a)) ⋅ (?x ⋆ ?x)"
unfolding Φ⇩P_def CMP.map_def
using assms B.VV.arr_char⇩S⇩b⇩C B.VV.cod_char⇩S⇩b⇩C P.FF_def by auto
also have "... = P 𝗂⇩B[a] ⋅ (CMP a a ⋅ (P a ⋆⇩B P a)) ⋅ (?x ⋆⇩B ?x)"
using assms x⇩S hcomp_char src_def trg_def by auto
also have "... = P 𝗂⇩B[a] ⋅⇩B (CMP a a ⋅⇩B (P a ⋆⇩B P a)) ⋅⇩B (?x ⋆⇩B ?x)"
proof -
have "«P 𝗂⇩B[a] ⋅ (CMP a a ⋅ (P a ⋆⇩B P a)) ⋅ (?x ⋆⇩B ?x) : P⇩0 a ⋆ P⇩0 a ⇒ P a»"
proof (intro comp_in_homI)
show "«?x ⋆⇩B ?x : P⇩0 a ⋆ P⇩0 a ⇒ P a ⋆ P a»"
proof -
have "«?x ⋆ ?x : P⇩0 a ⋆ P⇩0 a ⇒ P a ⋆ P a»"
using assms x⇩S * P⇩0_props P.map⇩0_simps(3)
by (intro hcomp_in_vhom) auto
moreover have "?x ⋆ ?x = ?x ⋆⇩B ?x"
using x⇩S * by (intro hcomp_eqI) auto
ultimately show ?thesis by simp
qed
show "«P a ⋆⇩B P a : P a ⋆ P a ⇒ P a ⋆ P a»"
using assms hcomp_eqI by fastforce
show "«CMP a a : P a ⋆ P a ⇒ P (a ⋆⇩B a)»"
using assms CMP_in_hom(2) by auto
show "«P 𝗂⇩B[a] : P (a ⋆⇩B a) ⇒ P a»"
using assms by auto
qed
thus ?thesis
by (intro comp_eqI hcomp_eqI) auto
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a] ⋅⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a)) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
(P a ⋆⇩B P a) ⋅⇩B
(?x ⋆⇩B ?x)"
using assms B.comp_assoc CMP_def by auto
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
(P a ⋆⇩B P a) ⋅⇩B
(?x ⋆⇩B ?x)"
using assms 1 B.whisker_left B.comp_assoc by fastforce
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((P a ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]))) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
using assms B.interchange B.comp_assoc by simp
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
((e a ⋆⇩B a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have "(P a ⋆⇩B P a) ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]))
= ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]))"
unfolding P_def
using assms B.comp_cod_arr [of "(e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
"(e a ⋆⇩B a ⋆⇩B d a) ⋆⇩B (e a ⋆⇩B a ⋆⇩B d a)"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
((e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have "(e a ⋆⇩B a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
= e a ⋆⇩B (a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B (a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
proof -
have "B.seq (a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) (a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
using assms by (elim B.objE, intro B.seqI) auto
thus ?thesis
using assms B.whisker_left by simp
qed
also have "... = e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"
using assms 1 B.whisker_left by fastforce
finally have "(e a ⋆⇩B a ⋆⇩B B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
= e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
(𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a)) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have "(e a ⋆⇩B a ⋆⇩B
(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B 𝖺⇩B[e a, a, d a ⋆⇩B P a]
= 𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
using assms 1 a_da P_def
B.assoc_naturality
[of "e a" a "(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"]
by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
(((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, d a ⋆⇩B P a]) ⋅⇩B
(e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B
𝖺⇩B[e a, a ⋆⇩B d a, P a] ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have 1: "B.ide (e a) ∧ B.ide a ∧ B.ide (d a) ∧ B.ide (P a)"
using assms ide_char⇩S⇩b⇩C P.preserves_ide by auto
have 2: "src⇩B (e a) = trg⇩B a ∧ src⇩B a = trg⇩B (d a) ∧ src⇩B (d a) = trg⇩B (P a)"
using assms by auto
have "((e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B 𝖺⇩B[e a, a ⋆⇩B d a, P a]) ⋅⇩B (𝖺⇩B[e a, a, d a] ⋆⇩B P a)
= 𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B 𝖺⇩B[e a ⋆⇩B a, d a, P a]"
using assms 1 2 B.pentagon B.comp_assoc by force
hence "(e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B 𝖺⇩B[e a, a ⋆⇩B d a, P a]
= (𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B 𝖺⇩B[e a ⋆⇩B a, d a, P a]) ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a)"
using assms 1 2
B.invert_side_of_triangle(2)
[of "𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B 𝖺⇩B[e a ⋆⇩B a, d a, P a]"
"(e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B 𝖺⇩B[e a, a ⋆⇩B d a, P a]"
"𝖺⇩B[e a, a, d a] ⋆⇩B P a"]
by fastforce
hence "𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B (𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a)
= 𝖺⇩B⇧-⇧1[e a, a, d a ⋆⇩B P a] ⋅⇩B (e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B
𝖺⇩B[e a, a ⋆⇩B d a, P a]"
using assms 1 2 B.invert_side_of_triangle(1) B.comp_assoc by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
(e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
((e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B
𝖺⇩B[e a, a ⋆⇩B d a, P a]) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have "B.arr (e a) ∧ B.arr a"
using assms by auto
moreover have "B.dom (e a) = e a ∧ B.dom a = a ∧ B.cod a = a ∧
B.cod (e a) = e a"
using assms by auto
moreover have "B.seq (B.inv (η a) ⋆⇩B a ⋆⇩B d a) 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"
using assms by (elim B.objE, intro B.seqI) auto
moreover have "src⇩B a =
trg⇩B ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
using assms a_da trg_vcomp by fastforce
moreover have "src⇩B a = a ∧ trg⇩B a = a"
using assms by auto
moreover have "B.dom ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
= d a ⋆⇩B e a ⋆⇩B a ⋆⇩B d a"
using assms a_da by fastforce
moreover have "B.cod ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
= a ⋆⇩B a ⋆⇩B d a"
using assms a_da by fastforce
ultimately have "((e a ⋆⇩B a) ⋆⇩B
(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, d a ⋆⇩B P a]
= 𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
(e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
unfolding P_def
using assms
B.assoc'_naturality
[of "e a" a "(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have 1: "B.ide (e a) ∧ B.ide a ∧ B.ide (d a) ∧ B.ide (P a)"
using assms ide_char⇩S⇩b⇩C P.preserves_ide by auto
have 2: "src⇩B (e a) = trg⇩B a ∧ src⇩B a = trg⇩B (d a) ∧ src⇩B (d a) = trg⇩B (P a)"
using assms by auto
have "((e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B 𝖺⇩B[e a, a ⋆⇩B d a, P a]) ⋅⇩B (𝖺⇩B[e a, a, d a] ⋆⇩B P a)
= 𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B 𝖺⇩B[e a ⋆⇩B a, d a, P a]"
using assms 1 2 B.pentagon B.comp_assoc by fastforce
hence "(e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B 𝖺⇩B[e a, a ⋆⇩B d a, P a]
= 𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B 𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a)"
using assms 1 2 P.preserves_ide B.comp_assoc
B.invert_side_of_triangle(2)
[of "𝖺⇩B[e a, a, d a ⋆⇩B P a] ⋅⇩B 𝖺⇩B[e a ⋆⇩B a, d a, P a]"
"(e a ⋆⇩B 𝖺⇩B[a, d a, P a]) ⋅⇩B 𝖺⇩B[e a, a ⋆⇩B d a, P a]"
"𝖺⇩B[e a, a, d a] ⋆⇩B P a"]
by force
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))"
proof -
have "(e a ⋆⇩B a ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, d a ⋆⇩B P a]
= 𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B ((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
proof -
have 1: "B.seq (B.inv (η a) ⋆⇩B a ⋆⇩B d a) 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"
using assms by (elim B.objE, intro B.seqI) auto
moreover have "src⇩B a =
trg⇩B ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])"
using a_da by force
moreover have "B.arr a ∧ B.dom a = a ∧ B.cod a = a ∧ src⇩B a = a ∧ trg⇩B a = a"
using assms by auto
moreover have "B.dom ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
= d a ⋆⇩B e a ⋆⇩B a ⋆⇩B d a"
using a_da by auto
moreover have "B.cod ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
= a ⋆⇩B a ⋆⇩B d a"
using a_da by auto
ultimately show ?thesis
unfolding P_def
using assms
B.assoc_naturality
[of "e a" a "(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]"]
by auto
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
((𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B
B.inv (ε a))) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))
= ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
(B.inv (ε a) ⋆⇩B B.inv (ε a))
= (((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
((e a ⋆⇩B d a) ⋆⇩B B.inv (ε a))) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
using assms B.comp_arr_dom B.comp_cod_arr B.comp_assoc
B.interchange [of "e a ⋆⇩B d a" "B.inv (ε a)" "B.inv (ε a)" "P⇩0 a"]
by simp
also have "... = ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B (e a ⋆⇩B d a) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B
B.inv (ε a)) ⋅⇩B (B.inv (ε a) ⋆⇩B P⇩0 a)"
using assms B.interchange by simp
also have "... = ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩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 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
(𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(((e a ⋆⇩B a) ⋆⇩B d a) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))) ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B
B.inv (ε a))
= (𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P a ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B
B.inv (ε a))"
proof -
have "B.seq 𝖺⇩B⇧-⇧1[e a, a, d a] (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using assms a_da by (elim B.objE, intro B.seqI) auto
moreover have "B.seq (P a) ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))"
using assms a_da P_def by auto
ultimately show ?thesis
using assms B.interchange by simp
qed
also have "... = (𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B
B.inv (ε a))"
proof -
have "P a ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) = (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)"
using assms x P_def B.comp_cod_arr by blast
moreover have "𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) = 𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a"
proof -
have "𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])
= B.inv ((e a ⋆⇩B 𝗅⇩B[d a]) ⋅⇩B 𝖺⇩B[e a, a, d a])"
proof -
have "B.inv ((e a ⋆⇩B 𝗅⇩B[d a]) ⋅⇩B 𝖺⇩B[e a, a, d a])
= B.inv 𝖺⇩B[e a, a, d a] ⋅⇩B B.inv (e a ⋆⇩B 𝗅⇩B[d a])"
proof -
have "B.seq (e a ⋆⇩B 𝗅⇩B[d a]) 𝖺⇩B[e a, a, d a]"
proof -
have "B.iso ((e a ⋆⇩B 𝗅⇩B[d a]) ⋅⇩B 𝖺⇩B[e a, a, d a])"
using assms by (elim B.objE, intro B.isos_compose) auto
thus ?thesis
by blast
qed
moreover have "B.iso 𝖺⇩B[e a, a, d a]"
using assms by force
ultimately show ?thesis
using assms B.inv_comp by auto
qed
also have "... = 𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using assms
by (elim B.objE) (simp add: assms)
finally show ?thesis by simp
qed
also have "... = B.inv (𝗋⇩B[e a] ⋆⇩B d a)"
using assms B.triangle [of "d a" "e a"] by simp
also have "... = 𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a"
using assms by simp
finally show ?thesis by blast
qed
ultimately show ?thesis by simp
qed
also have "... = (((e a ⋆⇩B a) ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[e a, a, d a]) ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B
(e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋅⇩B P⇩0 a"
using assms 0 B.comp_cod_arr B.comp_arr_dom
by (elim B.objE) auto
also have "... = ((e a ⋆⇩B a) ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B P⇩0 a"
using B.comp_assoc by simp
also have "... = (((e a ⋆⇩B a) ⋆⇩B d a) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a)"
proof -
have "B.seq ((e a ⋆⇩B a) ⋆⇩B d a) (𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]))"
using assms 0 a_da
by (elim B.objE, intro B.seqI) auto
moreover have "B.seq ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) (P⇩0 a)"
using assms 0
apply (intro B.seqI, auto simp add: B.obj_simps(5) P⇩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 "(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B P a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))
= (((e a ⋆⇩B a) ⋆⇩B d a) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a)"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
(((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "𝖺⇩B[e a ⋆⇩B a, d a, P a] ⋅⇩B
(((e a ⋆⇩B a) ⋆⇩B d a) ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))
= ((e a ⋆⇩B a) ⋆⇩B d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a]"
proof -
have "B.hseq (e a) a"
using assms by force
moreover have "src⇩B (d a) = trg⇩B ?x"
using assms B.trg_vcomp [of "e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]" "B.inv (ε a)"] by simp
moreover have "B.cod ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) = P a"
using assms x by blast
ultimately show ?thesis
using assms B.assoc_naturality [of "e a ⋆⇩B a" "d a" ?x] by auto
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B
(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a] ⋅⇩B
(d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "((e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a]) ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))
= (e a ⋆⇩B a) ⋆⇩B (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a] ⋅⇩B
(d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))"
proof -
have "B.ide (e a ⋆⇩B a)"
using assms by force
moreover have "B.seq ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a])
(d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))"
using a_da B.whisker_left by auto
ultimately show ?thesis
using assms B.whisker_left B.comp_assoc by fastforce
qed
thus ?thesis by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B
(B.inv (η a) ⋆⇩B d a) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B
(d a ⋆⇩B B.inv (ε a))) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a] ⋅⇩B
(d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))
= (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B (B.inv (η a) ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B
(d a ⋆⇩B B.inv (ε a))"
proof -
have "(B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a] ⋅⇩B
(d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))
= (B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B (𝖺⇩B⇧-⇧1[d a, e a, a ⋆⇩B d a] ⋅⇩B
(d a ⋆⇩B e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B (d a ⋆⇩B B.inv (ε a))"
using assms B.whisker_left B.comp_assoc by simp
also have "... = ((B.inv (η a) ⋆⇩B a ⋆⇩B d a) ⋅⇩B ((d a ⋆⇩B e a) ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B
𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B (d a ⋆⇩B B.inv (ε a))"
using assms B.assoc'_naturality [of "d a" "e a" "𝗅⇩B⇧-⇧1[d a]"] B.comp_assoc by simp
also have "... = (B.inv (η a) ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B
(d a ⋆⇩B B.inv (ε a))"
using assms B.interchange [of "B.inv (η a)" "d a ⋆⇩B e a" "a ⋆⇩B d a" "𝗅⇩B⇧-⇧1[d a]"]
B.comp_arr_dom B.comp_cod_arr
by simp
also have "... = (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B (B.inv (η a) ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B
(d a ⋆⇩B B.inv (ε a))"
using assms B.interchange [of a "B.inv (η a)" "𝗅⇩B⇧-⇧1[d a]" "d a"]
B.comp_arr_dom B.comp_cod_arr B.comp_assoc
by simp
finally show ?thesis by blast
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝗅⇩B⇧-⇧1[d a] ⋅⇩B 𝗋⇩B[d a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
(B.inv (ε a) ⋆⇩B P⇩0 a)"
proof -
have "(B.inv (η a) ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B (d a ⋆⇩B B.inv (ε a))
= 𝗅⇩B⇧-⇧1[d a] ⋅⇩B 𝗋⇩B[d a]"
proof -
have "(B.inv (η a) ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[d a, e a, d a] ⋅⇩B (d a ⋆⇩B B.inv (ε a))
= B.inv (η a ⋆⇩B d a) ⋅⇩B B.inv 𝖺⇩B[d a, e a, d a] ⋅⇩B B.inv (d a ⋆⇩B ε a)"
using assms by simp
also have "... = B.inv ((d a ⋆⇩B ε a) ⋅⇩B 𝖺⇩B[d a, e a, d a] ⋅⇩B (η a ⋆⇩B d a))"
proof -
have "B.iso ((d a ⋆⇩B ε a) ⋅⇩B 𝖺⇩B[d a, e a, d a] ⋅⇩B (η a ⋆⇩B d a))"
using assms by (intro B.isos_compose) auto
thus ?thesis
using assms B.inv_comp_left B.comp_assoc by auto
qed
also have "... = B.inv (𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a])"
using assms triangle_right by simp
also have "... = 𝗅⇩B⇧-⇧1[d a] ⋅⇩B 𝗋⇩B[d a]"
using assms B.inv_comp by simp
finally show ?thesis by blast
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝗅⇩B⇧-⇧1[d a] ⋅⇩B 𝗋⇩B[d a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
𝗋⇩B⇧-⇧1[e a ⋆⇩B d a]) ⋅⇩B
B.inv (ε a) ⋅⇩B
𝗂⇩B[P⇩0 a]"
proof -
have "𝗋⇩B[e a ⋆⇩B d a] ⋅⇩B (B.inv (ε a) ⋆⇩B P⇩0 a) = B.inv (ε a) ⋅⇩B 𝗂⇩B[P⇩0 a]"
using assms B.runit_naturality [of "B.inv (ε a)"] B.unitor_coincidence P⇩0_props
by simp
hence "B.inv (ε a) ⋆⇩B P⇩0 a = 𝗋⇩B⇧-⇧1[e a ⋆⇩B d a] ⋅⇩B B.inv (ε a) ⋅⇩B 𝗂⇩B[P⇩0 a]"
using assms P⇩0_props(5) B.invert_side_of_triangle(1) by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋅⇩B 𝗂⇩B[P⇩0 a]"
proof -
have P⇩0_a: "B.obj (P⇩0 a) ∧ B.arr (P⇩0 a)"
using assms a_da by fastforce
have i_a: "B.𝔯 a = 𝗂⇩B[a]"
using assms B.unitor_coincidence B.𝔯_ide_simp B.obj_simps
by (metis (no_types, lifting) B.objE)
have "P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝗅⇩B⇧-⇧1[d a] ⋅⇩B 𝗋⇩B[d a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
𝗋⇩B⇧-⇧1[e a ⋆⇩B d a]
= e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]"
proof -
have "P 𝗂⇩B[a] ⋅⇩B
(e a ⋆⇩B 𝖺⇩B⇧-⇧1[a, a, d a]) ⋅⇩B
(e a ⋆⇩B a ⋆⇩B 𝗅⇩B[a ⋆⇩B d a]) ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B⇧-⇧1[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
𝖺⇩B[e a, a, a ⋆⇩B a ⋆⇩B d a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝗅⇩B⇧-⇧1[d a] ⋅⇩B 𝗋⇩B[d a]) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, P⇩0 a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B P⇩0 a) ⋅⇩B
𝗋⇩B⇧-⇧1[e a ⋆⇩B d a]
= ⦃(❙⟨e a❙⟩ ❙⋆ ❙𝗋❙[❙⟨a❙⟩⇩0❙] ❙⋆ ❙⟨d a❙⟩) ❙⋅
(❙⟨e a❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨a❙⟩⇩0, ❙⟨a❙⟩⇩0, ❙⟨d a❙⟩❙]) ❙⋅
(❙⟨e a❙⟩ ❙⋆ ❙⟨a❙⟩⇩0 ❙⋆ ❙𝗅❙[❙⟨a❙⟩⇩0 ❙⋆ ❙⟨d a❙⟩❙]) ❙⋅
❙𝖺❙[❙⟨e a❙⟩, ❙⟨a❙⟩⇩0, ❙⟨a❙⟩⇩0 ❙⋆ ❙⟨a❙⟩⇩0 ❙⋆ ❙⟨d a❙⟩❙] ❙⋅
❙𝖺⇧-⇧1❙[❙⟨e a❙⟩, ❙⟨a❙⟩⇩0, ❙⟨a❙⟩⇩0 ❙⋆ ❙⟨a❙⟩⇩0 ❙⋆ ❙⟨d a❙⟩❙] ❙⋅
❙𝖺❙[❙⟨e a❙⟩, ❙⟨a❙⟩⇩0, ❙⟨a❙⟩⇩0 ❙⋆ ❙⟨a❙⟩⇩0 ❙⋆ ❙⟨d a❙⟩❙] ❙⋅
((❙⟨e a❙⟩ ❙⋆ ❙⟨a❙⟩⇩0) ❙⋆ (❙⟨a❙⟩⇩0 ❙⋆ ❙𝗅⇧-⇧1❙[❙⟨d a❙⟩❙]) ❙⋅ ❙𝗅⇧-⇧1❙[❙⟨d a❙⟩❙] ❙⋅ ❙𝗋❙[❙⟨d a❙⟩❙]) ❙⋅
❙𝖺❙[❙⟨e a❙⟩ ❙⋆ ❙⟨a❙⟩⇩0, ❙⟨d a❙⟩, ❙⟨P⇩0 a❙⟩⇩0❙] ❙⋅
(❙𝖺⇧-⇧1❙[❙⟨e a❙⟩, ❙⟨a❙⟩⇩0, ❙⟨d a❙⟩❙] ❙⋅ (❙⟨e a❙⟩ ❙⋆ ❙𝗅⇧-⇧1❙[❙⟨d a❙⟩❙]) ❙⋆ ❙⟨P⇩0 a❙⟩⇩0) ❙⋅
❙𝗋⇧-⇧1❙[❙⟨e a❙⟩ ❙⋆ ❙⟨d a❙⟩❙]⦄"
unfolding P_def
using assms B.obj_def [of a] P⇩0_a i_a B.α_def B.α'.map_ide_simp
B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C
B.𝔩_ide_simp B.𝔯_ide_simp
by (elim B.objE) simp
also have "... = ⦃❙⟨e a❙⟩ ❙⋆ ❙𝗅⇧-⇧1❙[❙⟨d a❙⟩❙]⦄"
using assms P⇩0_a by (intro EV.eval_eqI) auto
also have "... = e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]"
using assms P_def B.𝔩_ide_simp by simp
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = ?x ⋅ 𝗂⇩B[P.map⇩0 a]"
proof -
have "B.arr 𝗂⇩B[P.map⇩0 a] ∧ src⇩B 𝗂⇩B[P.map⇩0 a] ∈ Obj ∧ trg⇩B 𝗂⇩B[P.map⇩0 a] ∈ Obj ∧
P⇩0 a ⋅⇩B P⇩0 a ∈ Obj ∧ B.cod 𝗂⇩B[P.map⇩0 a] = P⇩0 a"
using assms P⇩0_props arr_char⇩S⇩b⇩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⇩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 𝖺 𝗂⇩B src trg V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
comp hcomp 𝖺 𝗂⇩B src trg
E Φ⇩E P Φ⇩P
..
sublocale EoP: composite_pseudofunctor
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B comp hcomp 𝖺 𝗂⇩B src trg V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
P Φ⇩P E Φ⇩E
..
sublocale I⇩C: identity_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ..
sublocale I⇩S: identity_pseudofunctor comp hcomp 𝖺 𝗂⇩B src trg ..
no_notation B.in_hom (‹«_ : _ →⇩B _»›)
abbreviation (input) unit⇩0
where "unit⇩0 ≡ e"
definition unit⇩1
where "unit⇩1 f = (e (trg⇩B f) ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
𝖺⇩B[e (trg⇩B f), f, src⇩B f] ⋅⇩B
((e (trg⇩B f) ⋆⇩B f) ⋆⇩B B.inv (η (src⇩B f))) ⋅⇩B
𝖺⇩B[e (trg⇩B f) ⋆⇩B f, d (src⇩B f), e (src⇩B f)] ⋅⇩B
(𝖺⇩B⇧-⇧1[e (trg⇩B f), f, d (src⇩B f)] ⋆⇩B e (src⇩B f))"
abbreviation (input) η⇩0
where "η⇩0 ≡ unit⇩0"
abbreviation (input) η⇩1
where "η⇩1 ≡ unit⇩1"
lemma unit⇩1_in_hom [intro]:
assumes "B.ide f"
shows "«η⇩1 f : src⇩B f →⇩B P⇩0 (trg⇩B f)»"
and "«η⇩1 f : (e (trg⇩B f) ⋆⇩B f ⋆⇩B d (src⇩B f)) ⋆⇩B e (src⇩B f) ⇒⇩B e (trg⇩B f) ⋆⇩B f»"
proof -
show "«η⇩1 f : (e (trg⇩B f) ⋆⇩B f ⋆⇩B d (src⇩B f)) ⋆⇩B e (src⇩B f) ⇒⇩B e (trg⇩B f) ⋆⇩B f»"
using assms unit⇩1_def by force
thus "«η⇩1 f : src⇩B f →⇩B P⇩0 (trg⇩B f)»"
using assms B.vconn_implies_hpar(1-2) by auto
qed
lemma unit⇩1_simps [simp]:
assumes "B.ide f"
shows "B.arr (η⇩1 f)"
and "src⇩B (η⇩1 f) = src⇩B f" and "trg⇩B (η⇩1 f) = P⇩0 (trg⇩B f)"
and "B.dom (η⇩1 f) = (e (trg⇩B f) ⋆⇩B f ⋆⇩B d (src⇩B f)) ⋆⇩B e (src⇩B f)"
and "B.cod (η⇩1 f) = e (trg⇩B f) ⋆⇩B f"
and "B.iso (η⇩1 f)"
using assms unit⇩1_in_hom
apply auto
unfolding unit⇩1_def
by (intro B.isos_compose) auto
lemma unit⇩1_in_hom⇩S [intro]:
assumes "ide f"
shows "«η⇩1 f : src f → P⇩0 (trg f)»"
and "«η⇩1 f : PoE.map f ⋆ e (src f) ⇒ e (trg f) ⋆ I⇩S.map f»"
using assms ide_char⇩S⇩b⇩C arr_char⇩S⇩b⇩C P⇩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⇩S⇩b⇩C
by auto
lemma unit⇩1_simps⇩S [simp]:
assumes "ide f"
shows "arr (η⇩1 f)"
and "src (η⇩1 f) = src f" and "trg (η⇩1 f) = P⇩0 (trg f)"
and "dom (η⇩1 f) = PoE.map f ⋆ e (src f)"
and "cod (η⇩1 f) = e (trg f) ⋆ I⇩S.map f"
and "iso (η⇩1 f)"
using assms unit⇩1_in_hom⇩S iso_char⇩S⇩b⇩C ide_char⇩S⇩b⇩C arr_char⇩S⇩b⇩C P.as_nat_iso.components_are_iso
by auto
sublocale unit: pseudonatural_equivalence comp hcomp 𝖺 𝗂⇩B src trg
comp hcomp 𝖺 𝗂⇩B src trg
I⇩S.map I⇩S.cmp ‹P ∘ E› PoE.cmp unit⇩0 unit⇩1
proof
show "⋀a. obj a ⟹ ide (e a)"
using obj_char ide_char⇩S⇩b⇩C arr_char⇩S⇩b⇩C P⇩0_props(1) by force
show "⋀a. obj a ⟹ «e a : src (I⇩S.map a) → src ((P ∘ E) a)»"
using emb.map_def
apply (intro in_hhomI)
apply auto
using obj_char by auto
show "⋀a. obj a ⟹ equivalence_map (e a)"
proof -
fix a
assume a: "obj a"
interpret Adj': equivalence_in_bicategory comp hcomp 𝖺 𝗂⇩B src trg
‹e a› ‹d a› ‹η a› ‹ε a›
using a by unfold_locales auto
show "equivalence_map (e a)"
using Adj'.equivalence_in_bicategory_axioms equivalence_map_def by auto
qed
show "⋀f. ide f ⟹ «η⇩1 f : PoE.map f ⋆ e (src f) ⇒ e (trg f) ⋆ I⇩S.map f»"
using unit⇩1_in_hom⇩S(2) by simp
show "⋀f. ide f ⟹ iso (η⇩1 f)"
by simp
show *: "⋀μ. arr μ ⟹
η⇩1 (cod μ) ⋅ (PoE.map μ ⋆ e (src μ)) = (e (trg μ) ⋆ I⇩S.map μ) ⋅ η⇩1 (dom μ)"
proof -
fix μ
assume μ: "arr μ"
have 1: "B.arr μ"
using μ arr_char⇩S⇩b⇩C by simp
let ?f = "B.dom μ"
let ?g = "B.cod μ"
let ?a = "src⇩B μ"
let ?b = "trg⇩B μ"
have "η⇩1 (cod μ) ⋅ (PoE.map μ ⋆ e (src μ))
= ((e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?g, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?g, d ?a] ⋆⇩B e ?a)) ⋅⇩B
((e ?b ⋆⇩B μ ⋆⇩B d ?a) ⋆⇩B e ?a)"
unfolding unit⇩1_def P_def emb.map_def hcomp_def src_def
using μ comp_char cod_simp arr_char⇩S⇩b⇩C P⇩0_props [of ?a] P⇩0_props [of ?b]
by auto
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?g, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?g, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?b ⋆⇩B μ ⋆⇩B d ?a) ⋆⇩B e ?a)"
using 1 B.comp_assoc by simp
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?g, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?g, d ?a] ⋅⇩B (e ?b ⋆⇩B μ ⋆⇩B d ?a) ⋆⇩B e ?a ⋅⇩B e ?a)"
using 1 B.interchange [of "𝖺⇩B⇧-⇧1[e ?b, ?g, d ?a]" "e ?b ⋆⇩B μ ⋆⇩B d ?a" "e ?a" "e ?a"]
by simp
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?g, d ?a, e ?a] ⋅⇩B
(((e ?b ⋆⇩B μ) ⋆⇩B d ?a) ⋅⇩B 𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a ⋅⇩B e ?a)"
using 1 B.assoc'_naturality [of "e ?b" μ "d ?a"] by simp
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B
(𝖺⇩B[e ?b ⋆⇩B ?g, d ?a, e ?a] ⋅⇩B
(((e ?b ⋆⇩B μ) ⋆⇩B d ?a) ⋆⇩B e ?a)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a)"
using 1 B.comp_assoc
B.interchange [of "(e ?b ⋆⇩B μ) ⋆⇩B d ?a" "𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a]" "e ?a" "e ?a" ]
by simp
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
(((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B
((e ?b ⋆⇩B μ) ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a)"
using 1 B.assoc_naturality [of "e ?b ⋆⇩B μ" "d ?a" "e ?a"] B.comp_assoc by simp
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B μ) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a)"
proof -
have "((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B ((e ?b ⋆⇩B μ) ⋆⇩B d ?a ⋆⇩B e ?a)
= (e ?b ⋆⇩B ?g) ⋅⇩B (e ?b ⋆⇩B μ) ⋆⇩B B.inv (η ?a) ⋅⇩B (d ?a ⋆⇩B e ?a)"
using 1 B.interchange by simp
also have "... = (e ?b ⋆⇩B μ) ⋆⇩B B.inv (η ?a)"
using 1 B.comp_arr_dom B.comp_cod_arr by simp
finally have "((e ?b ⋆⇩B ?g) ⋆⇩B B.inv (η ?a)) ⋅⇩B ((e ?b ⋆⇩B μ) ⋆⇩B d ?a ⋆⇩B e ?a)
= (e ?b ⋆⇩B μ) ⋆⇩B B.inv (η ?a)"
by simp
thus ?thesis by simp
qed
also have "... = (e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
(𝖺⇩B[e ?b, ?g, ?a] ⋅⇩B
((e ?b ⋆⇩B μ) ⋆⇩B ?a)) ⋅⇩B
((e ?b ⋆⇩B ?f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a)"
proof -
have "(e ?b ⋆⇩B μ) ⋆⇩B B.inv (η ?a) = (e ?b ⋆⇩B μ) ⋅⇩B (e ?b ⋆⇩B ?f) ⋆⇩B ?a ⋅⇩B B.inv (η ?a)"
using 1 B.comp_arr_dom B.comp_cod_arr by simp
also have "... = ((e ?b ⋆⇩B μ) ⋆⇩B ?a) ⋅⇩B ((e ?b ⋆⇩B ?f) ⋆⇩B B.inv (η ?a))"
using 1 B.interchange by simp
finally have "(e ?b ⋆⇩B μ) ⋆⇩B B.inv (η ?a)
= ((e ?b ⋆⇩B μ) ⋆⇩B ?a) ⋅⇩B ((e ?b ⋆⇩B ?f) ⋆⇩B B.inv (η ?a))"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B
(e ?b ⋆⇩B μ ⋆⇩B ?a)) ⋅⇩B
𝖺⇩B[e ?b, ?f, ?a] ⋅⇩B
((e ?b ⋆⇩B ?f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a)"
using 1 B.assoc_naturality [of "e ?b" μ "?a"] B.comp_assoc by simp
also have "... = (e ?b ⋆⇩B μ) ⋅⇩B
(e ?b ⋆⇩B 𝗋⇩B[?f]) ⋅⇩B
𝖺⇩B[e ?b, ?f, ?a] ⋅⇩B
((e ?b ⋆⇩B ?f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, ?f, d ?a] ⋆⇩B e ?a)"
proof -
have "(e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B (e ?b ⋆⇩B μ ⋆⇩B ?a) = e ?b ⋅⇩B e ?b ⋆⇩B 𝗋⇩B[?g] ⋅⇩B (μ ⋆⇩B ?a)"
using 1 B.interchange [of "e ?b" "e ?b" "𝗋⇩B[?g]" "μ ⋆⇩B ?a"] by simp
also have "... = e ?b ⋅⇩B e ?b ⋆⇩B μ ⋅⇩B 𝗋⇩B[?f]"
using 1 B.runit_naturality by simp
also have "... = (e ?b ⋆⇩B μ) ⋅⇩B (e ?b ⋆⇩B 𝗋⇩B[?f])"
using 1 B.interchange [of "e ?b" "e ?b" μ "𝗋⇩B[?f]"] by simp
finally have "(e ?b ⋆⇩B 𝗋⇩B[?g]) ⋅⇩B (e ?b ⋆⇩B μ ⋆⇩B ?a) = (e ?b ⋆⇩B μ) ⋅⇩B (e ?b ⋆⇩B 𝗋⇩B[?f])"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e (trg μ) ⋆ I⇩S.map μ) ⋅ η⇩1 (dom μ)"
proof -
have "arr ((e (trg μ) ⋆ I⇩S.map μ) ⋅ η⇩1 (dom μ))"
using μ by simp
hence 2: "arr ((e ?b ⋆⇩B μ) ⋅
((e ?b ⋆⇩B 𝗋⇩B[?f]) ⋅⇩B
𝖺⇩B[e ?b, ?f, ?a] ⋅⇩B
((e ?b ⋆⇩B ?f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B ?f, d ?a, e ?a] ⋅⇩B
(B.inv 𝖺⇩B[e ?b, ?f, d ?a] ⋆⇩B e ?a)))"
unfolding unit⇩1_def
using μ 1 arr_char⇩S⇩b⇩C P⇩0_props trg_def hcomp_def dom_char⇩S⇩b⇩C by simp
show ?thesis
unfolding unit⇩1_def
using μ 1 arr_char⇩S⇩b⇩C P⇩0_props trg_def hcomp_def dom_char⇩S⇩b⇩C
apply simp
using 2 comp_eqI by blast
qed
finally show "η⇩1 (cod μ) ⋅ (PoE.map μ ⋆ e (src μ))
= (e (trg μ) ⋆ I⇩S.map μ) ⋅ η⇩1 (dom μ)"
by simp
qed
show "⋀a. obj a ⟹ (e a ⋆ I⇩S.unit a) ⋅ 𝗋⇧-⇧1[e a] ⋅ 𝗅[e a] = unit⇩1 a ⋅ (PoE.unit a ⋆ e a)"
proof -
fix a
assume a: "obj a"
have 0: "B.obj a"
using a obj_char by blast
interpret adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ‹e a› ‹d a› ‹η a› ‹ε a›
using 0 chosen_adjoint_equivalence(1) [of a] by simp
have src: "src⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) = P⇩0 a"
using 0 B.src_vcomp [of "e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]" "B.inv (ε a)"] emb.map_def by simp
have trg: "trg⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) = P⇩0 a"
using 0 B.trg_vcomp [of "e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]" "B.inv (ε a)"] emb.map_def by simp
have 1: "arr a ∧ src⇩B a = a"
using a src_def by auto
have 2: "seq (P a) ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))"
proof -
have "B.arr (P a ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)))"
using a 0 P_def emb.map_def
by (elim B.objE, intro B.seqI) auto
moreover have "arr (P a)"
using 0 by auto
moreover have "arr ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))"
apply (unfold arr_char⇩S⇩b⇩C)
by (simp add: 0 P⇩0_props(6) src trg)
ultimately show ?thesis
using seq_char⇩S⇩b⇩C by simp
qed
have 3: "obj (P⇩0 a)"
using 0 P⇩0_props arr_char⇩S⇩b⇩C obj_char by blast
have 4: "B.seq (P a) ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a))"
using 2 seq_char⇩S⇩b⇩C by simp
have "η⇩1 a ⋅ (PoE.unit a ⋆ e a) = η⇩1 a ⋅ (P a ⋅ (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a)"
proof -
have "η⇩1 a ⋅ (PoE.unit a ⋆ e a) = η⇩1 a ⋅ (PoE.unit a ⋆⇩B e a)"
using 0 a hcomp_eqI by force
moreover have "PoE.unit a = P a ⋅ (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)"
using a obj_char PoE.unit_char' emb.unit_char' prj_unit_char by simp
ultimately show ?thesis by argo
qed
also have "... = η⇩1 a ⋅ (P a ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a)"
using 2 comp_simp by force
also have "... = η⇩1 a ⋅⇩B (P a ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a)"
proof -
have "arr (η⇩1 a)"
using a by auto
moreover have "arr (P a ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a)"
by (metis (no_types, lifting) a 0 2 B.src_vcomp B.vseq_implies_hpar(1)
arr_char⇩S⇩b⇩C arr_compI equivalence_data_in_hom(3) equivalence_data_simps⇩B(6)
hcomp_closed in_hom_char⇩S⇩b⇩C src)
moreover have "B.seq (η⇩1 a) (P a ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a)"
using 0 3 P_def obj_char
by (elim B.objE, intro B.seqI) auto
ultimately show ?thesis
using 2 comp_char by meson
qed
also have "... = (e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B
𝖺⇩B[e a, a, a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B B.inv (η a)) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, e a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B e a) ⋅⇩B
(((e a ⋆⇩B a ⋆⇩B d a) ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])) ⋅⇩B B.inv (ε a) ⋆⇩B e a)"
using 0 1 unit⇩1_def P_def emb.map⇩0_def emb.map_def B.comp_assoc B.obj_def' by simp
also have "... = (e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B
𝖺⇩B[e a, a, a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B B.inv (η a)) ⋅⇩B
𝖺⇩B[e a ⋆⇩B a, d a, e a] ⋅⇩B
((𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B e a) ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a))"
proof -
have "(e a ⋆⇩B a ⋆⇩B d a) ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) = e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]"
using 0 B.comp_cod_arr by simp
thus ?thesis by simp
qed
also have "... = (e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B
𝖺⇩B[e a, a, a] ⋅⇩B
((e a ⋆⇩B a) ⋆⇩B B.inv (η a)) ⋅⇩B
(𝖺⇩B[e a ⋆⇩B a, d a, e a] ⋅⇩B
((𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a) ⋆⇩B e a)) ⋅⇩B
(B.inv (ε a) ⋆⇩B e a)"
proof -
have "(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B e a) ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a)
= (𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B e a) ⋅⇩B
((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B e a) ⋅⇩B (B.inv (ε a) ⋆⇩B e a)"
using 0 B.whisker_right by simp
also have "... = ((𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B e a) ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B e a)) ⋅⇩B
(B.inv (ε a) ⋆⇩B e a)"
using B.comp_assoc by simp
also have "... = (𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋆⇩B e a) ⋅⇩B (B.inv (ε a) ⋆⇩B e a)"
proof -
have "B.seq 𝖺⇩B⇧-⇧1[e a, a, d a] (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using a 0 by (elim B.objE, intro B.seqI) auto
thus ?thesis
using 0 B.whisker_right by simp
qed
also have "... = ((𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a) ⋆⇩B e a) ⋅⇩B (B.inv (ε a) ⋆⇩B e a)"
proof -
have "𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) = 𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a"
proof -
have "e a ⋆⇩B 𝗅⇩B[d a] = (𝗋⇩B[e a] ⋆⇩B d a) ⋅⇩B 𝖺⇩B⇧-⇧1[e a, a, d a]"
using 0 1 B.triangle' [of "e a" "d a"] by simp
moreover have "B.hseq (e a) 𝗅⇩B[d a]"
using 0 by auto
moreover have "src⇩B (e a) = a"
using 0 by simp
ultimately have "(𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a) ⋅⇩B (e a ⋆⇩B 𝗅⇩B[d a]) = 𝖺⇩B⇧-⇧1[e a, a, d a]"
using 0 B.invert_side_of_triangle(1)
[of "e a ⋆⇩B 𝗅⇩B[d a]" "𝗋⇩B[e a] ⋆⇩B d a" "𝖺⇩B⇧-⇧1[e a, src⇩B (e a), d a]"]
by simp
moreover have "B.arr 𝖺⇩B⇧-⇧1[e a, a, d a]"
using 0 B.assoc'_in_hom [of "e a" a "d a"] by fastforce
ultimately have "𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a = 𝖺⇩B⇧-⇧1[e a, a, d a] ⋅⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using 0 B.invert_side_of_triangle(2)
[of "𝖺⇩B⇧-⇧1[e a, a, d a]" "𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a" "e a ⋆⇩B 𝗅⇩B[d a]"]
by simp
thus ?thesis by simp
qed
thus ?thesis by simp
qed
finally have "(𝖺⇩B⇧-⇧1[e a, a, d a] ⋆⇩B e a) ⋅⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋆⇩B e a) =
((𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a) ⋆⇩B e a) ⋅⇩B (B.inv (ε a) ⋆⇩B e a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B
𝖺⇩B[e a, a, a] ⋅⇩B
(((e a ⋆⇩B a) ⋆⇩B B.inv (η a)) ⋅⇩B
(𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a ⋆⇩B e a)) ⋅⇩B
𝖺⇩B[e a, d a, e a] ⋅⇩B
(B.inv (ε a) ⋆⇩B e a)"
using 0 B.assoc_naturality [of "𝗋⇩B⇧-⇧1[e a]" "d a" "e a"] B.comp_assoc by simp
also have "... = (e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B
𝖺⇩B[e a, a, a] ⋅⇩B
(𝗋⇩B⇧-⇧1[e a] ⋆⇩B a) ⋅⇩B
((e a ⋆⇩B B.inv (η a)) ⋅⇩B
𝖺⇩B[e a, d a, e a] ⋅⇩B
(B.inv (ε a) ⋆⇩B e a))"
proof -
have "((e a ⋆⇩B a) ⋆⇩B B.inv (η a)) ⋅⇩B (𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a ⋆⇩B e a)
= (e a ⋆⇩B a) ⋅⇩B 𝗋⇩B⇧-⇧1[e a] ⋆⇩B B.inv (η a) ⋅⇩B (d a ⋆⇩B e a)"
using 0 B.interchange [of "e a ⋆⇩B a" "𝗋⇩B⇧-⇧1[e a]" "B.inv (η a)" "d a ⋆⇩B e a"]
by force
also have "... = 𝗋⇩B⇧-⇧1[e a] ⋅⇩B e a ⋆⇩B a ⋅⇩B B.inv (η a)"
using 0 B.comp_arr_dom B.comp_cod_arr by simp
also have "... = (𝗋⇩B⇧-⇧1[e a] ⋆⇩B a) ⋅⇩B (e a ⋆⇩B B.inv (η a))"
using 0 B.interchange [of "𝗋⇩B⇧-⇧1[e a]" "e a" a "B.inv (η a)"] by fastforce
ultimately have "((e a ⋆⇩B a) ⋆⇩B B.inv (η a)) ⋅⇩B (𝗋⇩B⇧-⇧1[e a] ⋆⇩B d a ⋆⇩B e a)
= (𝗋⇩B⇧-⇧1[e a] ⋆⇩B a) ⋅⇩B (e a ⋆⇩B B.inv (η a))"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B
𝖺⇩B[e a, a, a] ⋅⇩B
(𝗋⇩B⇧-⇧1[e a] ⋆⇩B a)) ⋅⇩B
𝗋⇩B⇧-⇧1[e a] ⋅⇩B 𝗅⇩B[e a]"
by (metis B.comp_assoc adjoint_equivalence_in_bicategory_def
adjunction_in_bicategory.triangle_right dual_adjoint_equivalence)
also have "... = 𝗋⇩B⇧-⇧1[e a] ⋅⇩B 𝗅⇩B[e a]"
proof -
have "(e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B 𝖺⇩B[e a, a, a] ⋅⇩B (𝗋⇩B⇧-⇧1[e a] ⋆⇩B a) = e a ⋆⇩B a"
proof -
have "(e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B 𝖺⇩B[e a, a, a] ⋅⇩B (𝗋⇩B⇧-⇧1[e a] ⋆⇩B a)
= ((e a ⋆⇩B 𝗋⇩B[a]) ⋅⇩B 𝖺⇩B[e a, a, a]) ⋅⇩B (𝗋⇩B⇧-⇧1[e a] ⋆⇩B a)"
using B.comp_assoc by simp
also have "... = (𝗋⇩B[e a] ⋆⇩B a) ⋅⇩B (𝗋⇩B⇧-⇧1[e a] ⋆⇩B a)"
using 0 B.runit_char(2) [of "e a"] B.unitor_coincidence by simp
also have "... = e a ⋆⇩B a"
using 0 B.whisker_right [of a "𝗋⇩B[e a]" "𝗋⇩B⇧-⇧1[e a]"] B.comp_arr_inv' by auto
finally show ?thesis
by simp
qed
thus ?thesis
using 0 B.comp_cod_arr by simp
qed
also have "... = (e a ⋆ I⇩S.unit a) ⋅ 𝗋⇧-⇧1[e a] ⋅ 𝗅[e a]"
proof -
have "(e a ⋆ I⇩S.unit a) ⋅ 𝗋⇧-⇧1[e a] ⋅ 𝗅[e a] = (e a ⋆⇩B a) ⋅⇩B 𝗋⇩B⇧-⇧1[e a] ⋅⇩B 𝗅⇩B[e a]"
proof -
have 1: "arr 𝗅⇩B[e a] ∧ arr 𝗋⇩B⇧-⇧1[e a]"
using a lunit_simp runit'_simp lunit_simps runit'_simps by simp
moreover have "arr (𝗋⇩B⇧-⇧1[e a] ⋅⇩B 𝗅⇩B[e a])"
using 1 arr_char⇩S⇩b⇩C dom_char⇩S⇩b⇩C cod_char⇩S⇩b⇩C comp_char by auto
moreover have 2: "hseq (e a) (I⇩S.unit a)"
using a I⇩S.unit_char by (intro hseqI') auto
moreover have "B.seq (e a ⋆ I⇩S.unit a) (𝗋⇩B⇧-⇧1[e a] ⋅⇩B 𝗅⇩B[e a])"
using 0 2 hcomp_char arr_char⇩S⇩b⇩C I⇩S.unit_char'
apply (intro B.seqI)
apply auto[4]
using a by force
ultimately show ?thesis
using a comp_char I⇩S.unit_char' lunit_simp runit'_simp hcomp_char
by auto
qed
also have "... = 𝗋⇩B⇧-⇧1[e a] ⋅⇩B 𝗅⇩B[e a]"
using 0 B.comp_cod_arr by simp
finally show ?thesis by simp
qed
finally show "(e a ⋆ I⇩S.unit a) ⋅ 𝗋⇧-⇧1[e a] ⋅ 𝗅[e a] = η⇩1 a ⋅ (PoE.unit a ⋆ e a)"
by argo
qed
show "⋀f g. ⟦ide f; ide g; src g = trg f⟧ ⟹
(e (trg g) ⋆ I⇩S.cmp (g, f)) ⋅
𝖺[e (trg g), I⇩S.map g, I⇩S.map f] ⋅
(η⇩1 g ⋆ I⇩S.map f) ⋅
inv 𝖺[PoE.map g, e (src g), I⇩S.map f] ⋅
(PoE.map g ⋆ η⇩1 f) ⋅
𝖺[PoE.map g, PoE.map f, e (src f)]
= η⇩1 (g ⋆ f) ⋅ (PoE.cmp (g, f) ⋆ e (src f))"
proof -
fix f g
assume f: "ide f" and g: "ide g" and fg: "src g = trg f"
let ?a = "src⇩B f"
let ?b = "src⇩B g"
let ?c = "trg⇩B g"
have hseq: "ide f ∧ ide g ∧ arr f ∧ arr g ∧ src g = trg f"
using f g fg ide_char⇩S⇩b⇩C by auto
have hseq⇩B: "B.ide f ∧ B.ide g ∧ B.arr f ∧ B.arr g ∧ src⇩B g = trg⇩B f"
using f g fg ide_char⇩S⇩b⇩C src_def trg_def by auto
have 1: "?a = src f ∧ ?b = src g ∧ ?c = trg g"
using f g fg src_def trg_def by simp
have "(e (trg g) ⋆ I⇩S.cmp (g, f)) ⋅
𝖺[e (trg g), I⇩S.map g, I⇩S.map f] ⋅
(η⇩1 g ⋆ I⇩S.map f) ⋅
inv 𝖺[PoE.map g, e (src g), I⇩S.map f] ⋅
(PoE.map g ⋆ η⇩1 f) ⋅
𝖺[PoE.map g, PoE.map f, e (src f)]
= (e ?c ⋆ I⇩S.cmp (g, f)) ⋅
𝖺[e ?c, I⇩S.map g, I⇩S.map f] ⋅
(η⇩1 g ⋆ I⇩S.map f) ⋅
inv 𝖺[PoE.map g, e ?b, I⇩S.map f] ⋅
(PoE.map g ⋆ η⇩1 f) ⋅
𝖺[PoE.map g, PoE.map f, e ?a]"
using 1 by simp
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(e ?c ⋆ I⇩S.cmp (g, f)) ⋅
𝖺[e ?c, I⇩S.map g, I⇩S.map f] ⋅
(η⇩1 g ⋆ I⇩S.map f) ⋅
inv 𝖺[PoE.map g, e ?b, I⇩S.map f] ⋅
(PoE.map g ⋆ unit⇩1 f) ⋅
𝖺[PoE.map g, PoE.map f, e ?a]
= (e ?c ⋆ g ⋆ f) ⋅
𝖺[e ?c, g, f] ⋅
(η⇩1 g ⋆ f) ⋅
inv 𝖺[P (E g), e ?b, f] ⋅
(P (E g) ⋆ η⇩1 f) ⋅
𝖺[P (E g), P (E f), e ?a]"
using f g fg by simp
also have "... = (e ?c ⋆ g ⋆ f) ⋅⇩B
𝖺[e ?c, g, f] ⋅⇩B
(η⇩1 g ⋆ f) ⋅⇩B
inv 𝖺[P (E g), e ?b, f] ⋅⇩B
(P (E g) ⋆ η⇩1 f) ⋅⇩B
𝖺[P (E g), P (E f), e ?a]"
proof -
have "arr ..."
proof -
have "«... : (P (E g) ⋆ P (E f)) ⋆ e ?a ⇒ e ?c ⋆ g ⋆ f»"
proof (intro comp_in_homI)
show "«𝖺[P (E g), P (E f), e ?a] :
(P (E g) ⋆ P (E f)) ⋆ e ?a ⇒ P (E g) ⋆ P (E f) ⋆ e ?a»"
using f g fg VVV.arr_char⇩S⇩b⇩C VV.arr_char⇩S⇩b⇩C src_def obj_src assoc_in_hom
by force
show "«P (E g) ⋆ η⇩1 f : P (E g) ⋆ P (E f) ⋆ e ?a ⇒ P (E g) ⋆ e ?b ⋆ f»"
using f g fg 1 unit⇩1_in_hom⇩S [of f] vconn_implies_hpar(2)
by (intro hcomp_in_vhom) auto
show "«inv 𝖺[P (E g), e ?b, f] : P (E g) ⋆ e ?b ⋆ f ⇒ (P (E g) ⋆ e ?b) ⋆ f»"
proof -
have "«𝖺[P (E g), e ?b, f] : (P (E g) ⋆ e ?b) ⋆ f ⇒ P (E g) ⋆ e ?b ⋆ f»"
using f g fg 1 VVV.arr_char⇩S⇩b⇩C VV.arr_char⇩S⇩b⇩C assoc_in_hom by simp
moreover have "iso 𝖺[P (E g), e ?b, f]"
using f g fg 1 VVV.arr_char⇩S⇩b⇩C VV.arr_char⇩S⇩b⇩C iso_assoc by simp
ultimately show ?thesis
using inv_in_hom by simp
qed
show "«unit⇩1 g ⋆ f : (P (E g) ⋆ e ?b) ⋆ f ⇒ (e ?c ⋆ g) ⋆ f»"
using f g fg 1 unit⇩1_in_hom⇩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⇩S⇩b⇩C obj_trg by auto
show "«𝖺[e ?c, g, f] : (e ?c ⋆ g) ⋆ f ⇒ e ?c ⋆ g ⋆ f»"
using f g fg 1 2 3 VVV.arr_char⇩S⇩b⇩C VV.arr_char⇩S⇩b⇩C assoc_in_hom [of "e ?c" g f]
by simp
show "«e ?c ⋆ g ⋆ f : e ?c ⋆ g ⋆ f ⇒ e ?c ⋆ g ⋆ f»"
using f g fg 1 2 3
by (intro hcomp_in_vhom) auto
qed
thus ?thesis by blast
qed
thus ?thesis
using comp_eqI by blast
qed
also have "... = (e ?c ⋆ g ⋆ f) ⋅⇩B
𝖺[e ?c, g, f] ⋅⇩B
(η⇩1 g ⋆ f) ⋅⇩B
B.inv 𝖺[P (E g), e ?b, f] ⋅⇩B
(P (E g) ⋆ η⇩1 f) ⋅⇩B
𝖺[P (E g), P (E f), e ?a]"
proof -
have "iso 𝖺[P (E g), e ?b, f]"
using f g fg hseq hseq⇩B src_def trg_def emb.map⇩0_def emb.map_def arr_char⇩S⇩b⇩C
obj_char P⇩0_props(6)
by (intro iso_assoc) auto
thus ?thesis
using f g fg inv_char⇩S⇩b⇩C [of "𝖺[P (E g), e ?b, f]"] by simp
qed
also have "... = (e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B
𝖺[e ?c, g, f] ⋅⇩B
(η⇩1 g ⋆⇩B f) ⋅⇩B
B.inv 𝖺[P (E g), e ?b, f] ⋅⇩B
(P (E g) ⋆⇩B η⇩1 f) ⋅⇩B
𝖺[P (E g), P (E f), e ?a]"
proof -
have "arr (unit⇩0 ?c)"
using hseq obj_trg trg_def by auto
moreover have "arr (g ⋆⇩B f)"
using hseq hseq⇩B arr_char⇩S⇩b⇩C by auto
ultimately show ?thesis
unfolding hcomp_def
using f g fg hseq⇩B src_def trg_def P⇩0_props(1) emb.map⇩0_def emb.map_def
by simp
qed
also have "... = (e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g, f] ⋅⇩B
(η⇩1 g ⋆⇩B f) ⋅⇩B
B.inv 𝖺⇩B[P (E g), e ?b, f] ⋅⇩B
(P (E g) ⋆⇩B η⇩1 f) ⋅⇩B
𝖺⇩B[P (E g), P (E f), e ?a]"
proof -
have "arr (e ?a) ∧ arr (e ?b) ∧ arr (e ?c)"
using hseq hseq⇩B src_def trg_def obj_char src.preserves_arr trg.preserves_arr
by auto
thus ?thesis
using f g fg VVV.arr_char⇩S⇩b⇩C VV.arr_char⇩S⇩b⇩C src_def trg_def hseq⇩B
emb.map⇩0_def emb.map_def
by simp
qed
also have "... = (e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g, f] ⋅⇩B
(η⇩1 g ⋆⇩B f) ⋅⇩B
(B.inv 𝖺⇩B[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B η⇩1 f) ⋅⇩B
𝖺⇩B[P g, P f, e ?a])"
using f g fg 1 emb.map_def by simp
also have "... = (e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g, f] ⋅⇩B
(η⇩1 g ⋆⇩B f) ⋅⇩B
(𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B η⇩1 f) ⋅⇩B
𝖺⇩B[P g, P f, e ?a])"
using hseq⇩B P.preserves_ide ide_char⇩S⇩b⇩C B.assoc'_eq_inv_assoc [of "P g" "e ?b" f]
by simp
also have "... = (e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g, f] ⋅⇩B
((e ?c ⋆⇩B 𝗋⇩B[g]) ⋅⇩B
𝖺⇩B[e ?c, g, ?b] ⋅⇩B
((e ?c ⋆⇩B g) ⋆⇩B B.inv (η ?b)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b)
⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B (e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
𝖺⇩B[e ?b, f, ?a] ⋅⇩B
((e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
unfolding unit⇩1_def
using hseq⇩B B.comp_assoc by auto
also have "... = ((e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g, f]) ⋅⇩B
((e ?c ⋆⇩B 𝗋⇩B[g]) ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c, g, ?b] ⋆⇩B f) ⋅⇩B
(((e ?c ⋆⇩B g) ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B
(P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
using hseq⇩B ide_char⇩S⇩b⇩C P.preserves_ide B.whisker_left B.whisker_right B.comp_assoc
by auto
also have "... = 𝖺⇩B[e ?c, g, f] ⋅⇩B
((e ?c ⋆⇩B 𝗋⇩B[g]) ⋆⇩B f) ⋅⇩B
((𝖺⇩B[e ?c, g, ?b] ⋆⇩B f) ⋅⇩B
(((e ?c ⋆⇩B g) ⋆⇩B B.inv (η ?b)) ⋆⇩B f)) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B
(P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(e ?c ⋆⇩B g ⋆⇩B f) ⋅⇩B 𝖺⇩B[e ?c, g, f] = 𝖺⇩B[e ?c, g, f]"
using hseq⇩B B.comp_cod_arr by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = 𝖺⇩B[e ?c, g, f] ⋅⇩B
(((e ?c ⋆⇩B 𝗋⇩B[g]) ⋆⇩B f) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b)) ⋆⇩B f)) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B
(P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(𝖺⇩B[e ?c, g, ?b] ⋆⇩B f) ⋅⇩B (((e ?c ⋆⇩B g) ⋆⇩B B.inv (η ?b)) ⋆⇩B f)
= ((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B (𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f)"
proof -
have "(𝖺⇩B[e ?c, g, ?b] ⋆⇩B f) ⋅⇩B (((e ?c ⋆⇩B g) ⋆⇩B B.inv (η ?b)) ⋆⇩B f)
= 𝖺⇩B[e ?c, g, ?b] ⋅⇩B ((e ?c ⋆⇩B g) ⋆⇩B B.inv (η ?b)) ⋆⇩B f"
using hseq⇩B B.whisker_right by auto
also have "... = (e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b)) ⋅⇩B 𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f"
using hseq⇩B B.assoc_naturality [of "e ?c" g "B.inv (η ?b)"] by simp
also have "... = ((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f)"
using hseq⇩B B.whisker_right by auto
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[e ?c, g, f] ⋅⇩B
((e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B
(P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
using hseq⇩B B.whisker_left B.whisker_right B.comp_assoc by simp
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
((P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B
(P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
using hseq⇩B B.comp_assoc
B.assoc_naturality [of "e ?c" "𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))" f]
by fastforce
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
((P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f]) ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B f ⋆⇩B B.inv (η ?a))) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B (P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a))
= (P g ⋆⇩B e ?b ⋆⇩B f ⋆⇩B B.inv (η ?a)) ⋅⇩B (P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a])"
proof -
have "(P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a]) ⋅⇩B (P g ⋆⇩B (e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a))
= P g ⋆⇩B 𝖺⇩B[e ?b, f, ?a] ⋅⇩B ((e ?b ⋆⇩B f) ⋆⇩B B.inv (η ?a))"
using hseq⇩B ide_char⇩S⇩b⇩C P.preserves_ide B.whisker_left by auto
also have "... = P g ⋆⇩B (e ?b ⋆⇩B f ⋆⇩B B.inv (η ?a)) ⋅⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]"
using hseq⇩B B.assoc_naturality [of "e ?b" f "B.inv (η ?a)"] by auto
also have "... = (P g ⋆⇩B e ?b ⋆⇩B f ⋆⇩B B.inv (η ?a)) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a])"
using hseq⇩B ide_char⇩S⇩b⇩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 ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
(𝖺⇩B⇧-⇧1[P g, e ?b, f] ⋅⇩B
(P g ⋆⇩B e ?b ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
using hseq⇩B ide_char⇩S⇩b⇩C P.preserves_ide B.whisker_left B.comp_assoc by auto
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
(((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
((P g ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
using hseq⇩B B.comp_assoc
B.assoc'_naturality [of "P g" "e ?b" "𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"]
by simp
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
((P g ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
proof -
have "((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B f) ⋅⇩B
((P g ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋅⇩B (P g ⋆⇩B e ?b) ⋆⇩B
(f ⋅⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))"
proof -
have "B.seq (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) (P g ⋆⇩B e ?b)"
using hseq⇩B P.preserves_ide P_def src_def trg_def by auto
moreover have "B.seq f (𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))"
using hseq⇩B by simp
ultimately show ?thesis
using B.interchange
[of "𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b" "P g ⋆⇩B e ?b"
f "𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"]
by presburger
qed
also have "... = (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋅⇩B (e ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B e ?b)
⋆⇩B f ⋅⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩B B.whisker_right P_def by auto
also have "... = (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩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 ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
((𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B ((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b)
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= (𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a))"
proof -
have "(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= 𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋅⇩B (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b)
⋆⇩B f ⋅⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩B B.interchange by simp
also have "... = 𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋅⇩B (𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b)
⋆⇩B (𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.comp_cod_arr B.comp_arr_dom by simp
also have "... = (𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a))"
using hseq⇩B B.interchange by auto
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= 𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋅⇩B 𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b]
⋆⇩B f ⋅⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩B B.interchange by simp
also have "... = 𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋅⇩B 𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b]
⋆⇩B (𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.comp_cod_arr B.comp_arr_dom by simp
also have "... = (𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.interchange by auto
finally have "(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f) ⋅⇩B (𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b]
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= (𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
((e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= 𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B ((e ?c ⋆⇩B g ⋆⇩B d ?b ⋆⇩B e ?b) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b]
⋆⇩B (𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a))"
using hseq⇩B B.comp_arr_dom B.comp_cod_arr by simp
also have "... = 𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.interchange by simp
also have "... = (𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
using B.comp_assoc by simp
also have "... = (((e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a]) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.assoc_naturality [of "e ?c" "g ⋆⇩B d ?b ⋆⇩B e ?b"
"𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"]
by auto
also have "... = ((e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
using B.comp_assoc by simp
finally have "𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f] ⋅⇩B (𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b]
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))
= ((e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
have "(e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
((e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))))
= e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋅⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b)
⋆⇩B f ⋅⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩B B.whisker_left B.interchange by fastforce
also have "... = e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩B B.comp_arr_dom B.comp_cod_arr by auto
finally have "(e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋅⇩B
((e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))))
= e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
by simp
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, ?a] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "(e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, ?a] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)
= (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
(𝖺⇩B[e ?c, g ⋆⇩B f, ?a] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
using B.comp_assoc by simp
also have "... = ((e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
(e ?c ⋆⇩B (g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
using hseq⇩B B.assoc_naturality [of "e ?c" "g ⋆⇩B f" "B.inv (η ?a)"] B.comp_assoc
by simp
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a)) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
using hseq⇩B B.whisker_left B.comp_assoc by auto
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝗋⇩B[g] ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a)) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a)
= ((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝗋⇩B[g] ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a)"
using hseq⇩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 ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝗋⇩B[g] ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (((g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a))) ⋆⇩B e ?a)) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a)
= (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, ?b, f ⋆⇩B d ?a] ⋅⇩B (g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a)) ⋆⇩B e ?a"
using hseq⇩B B.whisker_right B.whisker_left by auto
also have "... = (e ?c ⋆⇩B
(((g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])
⋆⇩B e ?a"
using hseq⇩B B.assoc'_naturality [of g "B.inv (η ?b)" "f ⋆⇩B d ?a"] by auto
also have "... = ((e ?c ⋆⇩B (((g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a))) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a)"
using hseq⇩B B.whisker_right B.whisker_left by auto
finally have "((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a)
= ((e ?c ⋆⇩B (((g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a))) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a)) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
using hseq⇩B B.whisker_right B.whisker_left B.comp_assoc by auto
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a)) ⋆⇩B e ?a)) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a)
= (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a] ⋅⇩B
((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f ⋆⇩B d ?a)) ⋆⇩B e ?a"
using hseq⇩B B.whisker_left B.whisker_right by auto
also have "... = (e ?c ⋆⇩B
(((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a])
⋆⇩B e ?a"
using hseq⇩B B.assoc'_naturality [of "𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))" f "d ?a"]
by auto
also have "... = ((e ?c ⋆⇩B (((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a))
⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a)"
using hseq⇩B B.whisker_left B.whisker_right by auto
finally have "((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a)
= ((e ?c ⋆⇩B (((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a)) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a) ⋆⇩B e ?a)) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a)) ⋆⇩B e ?a)
= 𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋅⇩B
(e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋆⇩B d ?a)) ⋆⇩B e ?a"
using hseq⇩B B.whisker_right by auto
also have "... = ((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a"
using hseq⇩B B.assoc'_naturality [of "e ?c" "(𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f" "d ?a"]
by fastforce
also have "... = (((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a)
⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a)"
using hseq⇩B B.whisker_right by auto
finally have "(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B (((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a)) ⋆⇩B e ?a)
= (((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a)"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
(𝖺⇩B[e ?c, g ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a) ⋆⇩B e ?a)
= ((e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f)) ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a]"
using hseq⇩B
B.assoc_naturality [of "(e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f))"
"d ?a" "e ?a"]
by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))) ⋅⇩B
(e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[e ?c, (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
using hseq⇩B B.comp_assoc
B.assoc_naturality [of "e ?c" "(𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f"
"d ?a ⋆⇩B e ?a"]
by force
also have "... = (e ?c ⋆⇩B (g ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B[g, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(e ?c ⋆⇩B ((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))) ⋆⇩B f) ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B[e ?c, (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B B.inv (η ?a))
= (e ?c ⋆⇩B (g ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(e ?c ⋆⇩B 𝖺⇩B[g, f, d ?a ⋆⇩B e ?a])"
using hseq⇩B B.runit_hcomp B.whisker_left B.comp_assoc
B.assoc_naturality [of g f "B.inv (η ?a)"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((e ?c ⋆⇩B (g ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a))) ⋅⇩B
(e ?c ⋆⇩B 𝖺⇩B[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
𝖺⇩B[e ?c, (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "(e ?c ⋆⇩B 𝖺⇩B[g, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋆⇩B d ?a ⋆⇩B e ?a)
= e ?c ⋆⇩B 𝖺⇩B[g, f, d ?a ⋆⇩B e ?a] ⋅⇩B
((𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.whisker_left by auto
also have "... = e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a ⋆⇩B e ?a]"
using hseq⇩B B.assoc_naturality [of "𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))" f "d ?a ⋆⇩B e ?a"]
by auto
also have "... = (e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
(e ?c ⋆⇩B 𝖺⇩B[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a ⋆⇩B e ?a])"
using hseq⇩B B.whisker_left by auto
finally have "(e ?c ⋆⇩B 𝖺⇩B[g, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f) ⋆⇩B d ?a ⋆⇩B e ?a)
= (e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
(e ?c ⋆⇩B 𝖺⇩B[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a ⋆⇩B e ?a])"
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B
𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
(e ?c ⋆⇩B 𝖺⇩B[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
𝖺⇩B[e ?c, (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "(e ?c ⋆⇩B (g ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a))
= e ?c ⋆⇩B (g ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
(𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.whisker_left by auto
also have "... = e ?c ⋆⇩B g ⋅⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))
⋆⇩B (𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)"
using hseq⇩B B.interchange by auto
also have "... = e ?c ⋆⇩B (g ⋅⇩B 𝗋⇩B[g]) ⋅⇩B (g ⋆⇩B B.inv (η ?b))
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)) ⋅⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)"
using B.comp_assoc by simp
also have "... = e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
using hseq⇩B B.comp_arr_dom B.comp_cod_arr by auto
finally have "(e ?c ⋆⇩B (g ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a)))) ⋅⇩B
(e ?c ⋆⇩B (𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a))
= e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b)) ⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))"
by simp
thus ?thesis
by simp
qed
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g] ⋅⇩B (g ⋆⇩B B.inv (η ?b))
⋆⇩B 𝗋⇩B[f] ⋅⇩B (f ⋆⇩B B.inv (η ?a))) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
proof -
let ?LHS = "(e ?c ⋆⇩B 𝖺⇩B[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
𝖺⇩B[e ?c, (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a ⋆⇩B e ?a] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B (g ⋆⇩B d ?b ⋆⇩B e ?b) ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, (g ⋆⇩B (d ?b ⋆⇩B e ?b)) ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g ⋆⇩B d ?b ⋆⇩B e ?b, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a])) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
let ?RHS = "𝖺⇩B[e ?c, g ⋆⇩B d ?b ⋆⇩B e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, e ?b] ⋆⇩B f ⋆⇩B d ?a ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B e ?b) ⋆⇩B (f ⋆⇩B d ?a ⋆⇩B e ?a)) ⋅⇩B
𝖺⇩B⇧-⇧1[P g, e ?b, f ⋆⇩B d ?a ⋆⇩B e ?a] ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b, f, d ?a ⋆⇩B e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B[e ?b ⋆⇩B f, d ?a, e ?a]) ⋅⇩B
(P g ⋆⇩B 𝖺⇩B⇧-⇧1[e ?b, f, d ?a] ⋆⇩B e ?a) ⋅⇩B
𝖺⇩B[P g, P f, e ?a]"
have "?LHS = ?RHS"
proof -
let ?LHSt = "(❙⟨e ?c❙⟩ ❙⋆ ❙𝖺❙[❙⟨g❙⟩ ❙⋆ ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩, ❙⟨f❙⟩, ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩❙]) ❙⋅
❙𝖺❙[❙⟨e ?c❙⟩, (❙⟨g❙⟩ ❙⋆ ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩) ❙⋆ ❙⟨f❙⟩, ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩❙] ❙⋅
❙𝖺❙[❙⟨e ?c❙⟩ ❙⋆ (❙⟨g❙⟩ ❙⋆ ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩) ❙⋆ ❙⟨f❙⟩, ❙⟨d ?a❙⟩, ❙⟨e ?a❙⟩❙] ❙⋅
(❙𝖺⇧-⇧1❙[❙⟨e ?c❙⟩, (❙⟨g❙⟩ ❙⋆ (❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩)) ❙⋆ ❙⟨f❙⟩, ❙⟨d ?a❙⟩❙] ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
((❙⟨e ?c❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨g❙⟩ ❙⋆ ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩, ❙⟨f❙⟩, ❙⟨d ?a❙⟩❙]) ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
(((❙⟨e ?c❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨g❙⟩, ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩, ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙])) ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
((❙⟨e ?c❙⟩ ❙⋆ ❙⟨g❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨d ?b❙⟩, ❙⟨e ?b❙⟩, ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙]) ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
(❙𝖺❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩ ❙⋆ ❙P❙R❙J f❙] ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
(❙𝖺❙[❙⟨e ?c❙⟩ ❙⋆ ❙⟨g❙⟩, ❙⟨d ?b❙⟩, ❙P❙R❙J f❙] ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
((❙𝖺⇧-⇧1❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩ ❙] ❙⋆ ❙P❙R❙J f) ❙⋆ ❙⟨e ?a❙⟩)"
let ?RHSt = "❙𝖺❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩ ❙⋆ ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩, ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩❙] ❙⋅
(❙𝖺❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩❙] ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
(❙𝖺❙[❙⟨e ?c❙⟩ ❙⋆ ❙⟨g❙⟩, ❙⟨d ?b❙⟩, ❙⟨e ?b❙⟩❙] ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
((❙𝖺⇧-⇧1❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩❙] ❙⋆ ❙⟨e ?b❙⟩) ❙⋆ ( ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩)) ❙⋅
❙𝖺⇧-⇧1❙[❙P❙R❙J g, ❙⟨e ?b❙⟩, ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩❙] ❙⋅
(❙P❙R❙J g ❙⋆ ❙𝖺❙[❙⟨e ?b❙⟩, ❙⟨f❙⟩, ❙⟨d ?a❙⟩ ❙⋆ ❙⟨e ?a❙⟩❙]) ❙⋅
(❙P❙R❙J g ❙⋆ ❙𝖺❙[❙⟨e ?b❙⟩ ❙⋆ ❙⟨f❙⟩, ❙⟨d ?a❙⟩, ❙⟨e ?a❙⟩❙]) ❙⋅
(❙P❙R❙J g ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨e ?b❙⟩, ❙⟨f❙⟩, ❙⟨d ?a❙⟩❙] ❙⋆ ❙⟨e ?a❙⟩) ❙⋅
❙𝖺❙[❙P❙R❙J g, ❙P❙R❙J f, ❙⟨e ?a❙⟩❙]"
have "?LHS = ⦃?LHSt⦄"
using f g fg hseq⇩B B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def
by auto
also have "... = ⦃?RHSt⦄"
using hseq⇩B by (intro EV.eval_eqI, auto)
also have "... = ?RHS"
using f g fg hseq⇩B B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def
by auto
finally show ?thesis by blast
qed
thus ?thesis using hseq⇩B by auto
qed
finally show ?thesis by simp
qed
also have "... = η⇩1 (g ⋆ f) ⋅ (PoE.cmp (g, f) ⋆ e (src f))"
proof -
have "η⇩1 (g ⋆ f) ⋅ (PoE.cmp (g, f) ⋆ e (src f))
= η⇩1 (g ⋆ f) ⋅ (P (E (g ⋆ f)) ⋅ P (Φ⇩E (g, f)) ⋅ Φ⇩P (E g, E f) ⋆ e (src f))"
using f g fg PoE.cmp_def VV.arr_char⇩S⇩b⇩C VV.dom_char⇩S⇩b⇩C by simp
also have "... = η⇩1 (g ⋆ f) ⋅ (P (E (g ⋆ f)) ⋅ P (g ⋆⇩B f) ⋅ Φ⇩P (E g, E f) ⋆ e (src f))"
using f g fg emb.cmp_def VV.arr_char⇩S⇩b⇩C by simp
also have "... = η⇩1 (g ⋆ f) ⋅ (P (E (g ⋆ f)) ⋅ Φ⇩P (E g, E f) ⋆ e (src f))"
using f g fg hseq⇩B comp_cod_arr emb.map_def hseq_char' prj.cmp_simps'(1,5)
by auto
also have "... = η⇩1 (g ⋆ f) ⋅ (P (g ⋆⇩B f) ⋅ Φ⇩P (g, f) ⋆ e (src f))"
using hseq⇩B hseq emb.map_def hcomp_char hseqI' by auto
also have "... = η⇩1 (g ⋆ f) ⋅ (Φ⇩P (g, f) ⋅ (P g ⋆ P f) ⋆ e (src f))"
using hseq⇩B B.VV.arr_char⇩S⇩b⇩C B.VV.dom_char⇩S⇩b⇩C B.VV.cod_char⇩S⇩b⇩C Φ⇩P.naturality [of "(g, f)"]
P.FF_def
by auto
also have "... = η⇩1 (g ⋆ f) ⋅ (Φ⇩P (g, f) ⋆ e (src f))"
proof -
have "Φ⇩P (g, f) ⋅ (P g ⋆ P f) = Φ⇩P (g, f)"
using f g fg comp_arr_dom [of "Φ⇩P (g, f)" "P g ⋆ P f"] VV.arr_char⇩S⇩b⇩C
by (metis (no_types, lifting) prj.cmp_simps'(1) Φ⇩P_simps(4) hseq⇩B)
thus ?thesis by simp
qed
also have "... = η⇩1 (g ⋆ f) ⋅⇩B (Φ⇩P (g, f) ⋆ e (src f))"
proof -
have "seq (η⇩1 (g ⋆ f)) (Φ⇩P (g, f) ⋆ e (src f))"
proof -
have "cod (Φ⇩P (g, f)) = P (g ⋆ f)"
using f g fg Φ⇩P_simps(5) [of g f] VV.arr_char⇩S⇩b⇩C
by (metis (no_types, lifting) hcomp_eqI hseq hseq⇩B hseqI')
moreover have "hseq (Φ⇩P (g, f)) (e (src f))"
using f g fg hseq⇩B 1 B.VV.arr_char⇩S⇩b⇩C by auto
ultimately show ?thesis
using f g fg hseq⇩B 1 B.VV.arr_char⇩S⇩b⇩C P⇩0_props prj.cmp_simps emb.map_def
Φ⇩P_simps(5) [of g f] hcomp_eqI B.VV.dom_char⇩S⇩b⇩C B.VV.cod_char⇩S⇩b⇩C P.FF_def
by auto
qed
thus ?thesis
using comp_eqI by simp
qed
also have "... = η⇩1 (g ⋆⇩B f) ⋅⇩B (Φ⇩P (g, f) ⋆⇩B e (src f))"
using f g fg hseq⇩B 1 B.VV.arr_char⇩S⇩b⇩C hcomp_eqI by simp
also have "... = ((e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, ?a] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a)) ⋅⇩B
(Φ⇩P (g, f) ⋆⇩B e ?a)"
unfolding unit⇩1_def
using hseq⇩B 1 comp_char by simp
also have "... = ((e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, ?a] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a)) ⋅⇩B
(((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B
𝗅⇩B[f ⋆⇩B d ?a] ⋅⇩B
(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f))
⋆⇩B e ?a)"
proof -
have "B.VV.ide (g, f)"
using f g fg ide_char⇩S⇩b⇩C src_def trg_def B.VV.ide_char⇩S⇩b⇩C by auto
hence "Φ⇩P (g, f) = CMP g f"
unfolding Φ⇩P_def
using f g fg ide_char⇩S⇩b⇩C CMP.map_simp_ide by simp
also have "... = (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a] ⋅⇩B
(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f)"
using hseq⇩B CMP_def by auto
finally have "Φ⇩P (g, f) = (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a] ⋅⇩B
(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f)"
by blast
thus ?thesis by simp
qed
also have "... = ((e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, src f] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a)) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f)
⋆⇩B e ?a)"
using f g fg hseq⇩B src_def trg_def B.whisker_left B.comp_assoc by simp
also have "... = (e ?c ⋆⇩B 𝗋⇩B[g ⋆⇩B f]) ⋅⇩B
𝖺⇩B[e ?c, g ⋆⇩B f, src f] ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B f) ⋆⇩B B.inv (η ?a)) ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g ⋆⇩B f, d ?a, e ?a] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g ⋆⇩B f, d ?a] ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
proof -
have "B.arr ((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f))"
using hseq⇩B ide_char⇩S⇩b⇩C P.preserves_ide P_def by auto
hence "(e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f)
⋆⇩B e ?a
= ((e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋆⇩B e ?a) ⋅⇩B
((e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋆⇩B e ?a) ⋅⇩B
(𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋆⇩B e ?a) ⋅⇩B
((𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋆⇩B e ?a)"
using hseq⇩B B.whisker_right by fastforce
thus ?thesis
using B.comp_assoc by simp
qed
finally show ?thesis
using f ide_char⇩S⇩b⇩C src_def by simp
qed
finally show "(e (trg g) ⋆ I⇩S.cmp (g, f)) ⋅
𝖺[e (trg g), I⇩S.map g, I⇩S.map f] ⋅
(η⇩1 g ⋆ I⇩S.map f) ⋅
inv 𝖺[PoE.map g, e (src g), I⇩S.map f] ⋅
(PoE.map g ⋆ η⇩1 f) ⋅
𝖺[PoE.map g, PoE.map f, e (src f)]
= η⇩1 (g ⋆ f) ⋅ (PoE.cmp (g, f) ⋆ e (src f))"
by blast
qed
qed
abbreviation (input) counit⇩0
where "counit⇩0 ≡ d"
definition counit⇩1
where "counit⇩1 f = 𝖺⇩B[d (trg⇩B f), e (trg⇩B f), f ⋆⇩B d (src⇩B f)] ⋅⇩B
(η (trg⇩B f) ⋆⇩B f ⋆⇩B d (src⇩B f)) ⋅⇩B
𝗅⇩B⇧-⇧1[f ⋆⇩B d (src⇩B f)]"
abbreviation (input) ε⇩0
where "ε⇩0 ≡ counit⇩0"
abbreviation (input) ε⇩1
where "ε⇩1 ≡ counit⇩1"
lemma counit⇩1_in_hom [intro]:
assumes "B.ide f"
shows "«ε⇩1 f : f ⋆⇩B d (src⇩B f) ⇒⇩B d (trg⇩B f) ⋆⇩B e (trg⇩B f) ⋆⇩B f ⋆⇩B d (src⇩B f)»"
using assms B.iso_is_arr
by (unfold counit⇩1_def, intro B.comp_in_homI' B.seqI B.hseqI' B.hcomp_in_vhom) auto
lemma counit⇩1_simps [simp]:
assumes "B.ide f"
shows "B.arr (ε⇩1 f)"
and "src⇩B (ε⇩1 f) = P⇩0 (src⇩B f)" and "trg⇩B (ε⇩1 f) = trg⇩B f"
and "B.dom (ε⇩1 f) = f ⋆⇩B d (src⇩B f)"
and "B.cod (ε⇩1 f) = d (trg⇩B f) ⋆⇩B e (trg⇩B f) ⋆⇩B f ⋆⇩B d (src⇩B f)"
and "B.iso (ε⇩1 f)"
using assms counit⇩1_in_hom
apply auto
using B.vconn_implies_hpar(1)
apply fastforce
using B.vconn_implies_hpar(2)
apply fastforce
unfolding counit⇩1_def
apply (intro B.isos_compose)
by auto
lemma technical:
assumes "B.ide f" and "B.ide g" and "src⇩B g = trg⇩B f"
shows "(ε⇩1 g ⋆⇩B P f) ⋅⇩B 𝖺⇩B⇧-⇧1[g, d (src⇩B g), P f] ⋅⇩B (g ⋆⇩B ε⇩1 f)
= (𝖺⇩B[d (trg⇩B g), e (trg⇩B g), g ⋆⇩B d (src⇩B g)] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d (trg⇩B g) ⋆⇩B e (trg⇩B g), g, d (src⇩B g)] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d (trg⇩B g) ⋆⇩B e (trg⇩B g)) ⋆⇩B g, d (src⇩B g), P f] ⋅⇩B
(((d (trg⇩B g) ⋆⇩B e (trg⇩B g)) ⋆⇩B g) ⋆⇩B d (src⇩B g) ⋆⇩B P f) ⋅⇩B
(((d (trg⇩B g) ⋆⇩B e (trg⇩B g)) ⋆⇩B g)
⋆⇩B 𝖺⇩B[d (src⇩B g), e (src⇩B g), f ⋆⇩B d (src⇩B f)]) ⋅⇩B
((η (trg⇩B g) ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B (η (src⇩B g) ⋆⇩B f ⋆⇩B d (src⇩B f))) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[src⇩B g, f, d (src⇩B f)] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d (src⇩B f)))"
proof -
let ?a = "src⇩B f"
let ?b = "src⇩B g"
let ?c = "trg⇩B g"
have "(ε⇩1 g ⋆⇩B P f) ⋅⇩B 𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B (g ⋆⇩B ε⇩1 f)
= (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋅⇩B
(η ?c ⋆⇩B g ⋆⇩B d ?b) ⋅⇩B
𝗅⇩B⇧-⇧1[g ⋆⇩B d ?b]
⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
using assms counit⇩1_def by simp
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
((η ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
(𝗅⇩B⇧-⇧1[g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
proof -
have "𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋅⇩B
(η ?c ⋆⇩B g ⋆⇩B d ?b) ⋅⇩B
𝗅⇩B⇧-⇧1[g ⋆⇩B d ?b]
⋆⇩B P f
= (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
((η ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
(𝗅⇩B⇧-⇧1[g ⋆⇩B d ?b] ⋆⇩B P f)"
using assms B.iso_is_arr B.whisker_right P_def by auto
moreover have "g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a]
= (g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
using assms B.iso_is_arr B.whisker_left by simp
ultimately show ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
((η ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[?c, g, d ?b] ⋅⇩B (𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
using assms B.lunit_hcomp [of g "d ?b"] B.comp_assoc by simp
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(((η ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[?c, g, d ?b] ⋆⇩B P f)) ⋅⇩B
((𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
using assms B.whisker_right P_def B.comp_assoc by simp
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
(((η ?c ⋆⇩B g) ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
(((𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f]) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
proof -
have "((η ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B (𝖺⇩B[?c, g, d ?b] ⋆⇩B P f) =
(η ?c ⋆⇩B g ⋆⇩B d ?b) ⋅⇩B 𝖺⇩B[?c, g, d ?b] ⋆⇩B P f"
using assms B.whisker_right P_def B.iso_is_arr by simp
also have "... = 𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋅⇩B ((η ?c ⋆⇩B g) ⋆⇩B d ?b) ⋆⇩B P f"
using assms B.assoc_naturality [of "η ?c" g "d ?b"] B.iso_is_arr by simp
also have "... = (𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
(((η ?c ⋆⇩B g) ⋆⇩B d ?b) ⋆⇩B P f)"
using assms B.whisker_right P_def B.iso_is_arr by simp
finally have "((η ?c ⋆⇩B g ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B (𝖺⇩B[?c, g, d ?b] ⋆⇩B P f) =
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
(((η ?c ⋆⇩B g) ⋆⇩B d ?b) ⋆⇩B P f)"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
((((η ?c ⋆⇩B g) ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[?c ⋆⇩B g, d ?b, P f]) ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
proof -
have "((𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B 𝖺⇩B⇧-⇧1[g, d ?b, P f] =
𝖺⇩B⇧-⇧1[?c ⋆⇩B g, d ?b, P f] ⋅⇩B (𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b ⋆⇩B P f)"
using assms B.assoc'_naturality [of "𝗅⇩B⇧-⇧1[g]" "d ?b" "P f"] by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((η ?c ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b ⋆⇩B P f)) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
proof -
have "(((η ?c ⋆⇩B g) ⋆⇩B d ?b) ⋆⇩B P f) ⋅⇩B 𝖺⇩B⇧-⇧1[?c ⋆⇩B g, d ?b, P f] =
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B ((η ?c ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f)"
using assms B.assoc'_naturality [of "η ?c ⋆⇩B g" "d ?b" "P f"] B.iso_is_arr
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a])) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
using assms B.whisker_right B.iso_is_arr P_def B.comp_assoc by simp
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])"
proof -
have "((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a])
= (((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a])"
using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def
B.interchange [of "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]" g]
B.interchange [of "(d ?c ⋆⇩B e ?c) ⋆⇩B g" "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a]))"
proof -
have "((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B (g ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a)
= (((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a)"
using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def
B.interchange [of "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]" g]
B.interchange [of "(d ?c ⋆⇩B e ?c) ⋆⇩B g" "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a))"
proof -
have "((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])
= (η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a]"
using assms B.comp_arr_dom B.iso_is_arr P_def
B.interchange [of "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]" g "η ?b ⋆⇩B f ⋆⇩B d ?a" "𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a]"]
by simp
also have "... = (η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B
(η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[?b, f, d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)"
using assms B.lunit_hcomp by simp
also have "... = ((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a))"
using assms B.comp_arr_dom B.iso_is_arr P_def
B.interchange [of "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]" g]
by simp
finally have "((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (g ⋆⇩B 𝗅⇩B⇧-⇧1[f ⋆⇩B d ?a])
= ((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a))"
by blast
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
sublocale counit: pseudonatural_equivalence
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹E ∘ P› EoP.cmp I⇩C.map I⇩C.cmp counit⇩0 counit⇩1
proof
show "⋀a. B.obj a ⟹ B.ide (d a)"
by simp
show "⋀a. B.obj a ⟹ B.equivalence_map (d a)"
proof -
fix a
assume a: "B.obj a"
interpret Adj: adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
‹e a› ‹d a› ‹η a› ‹ε a›
using a chosen_adjoint_equivalence by simp
show "B.equivalence_map (d a)"
using Adj.equivalence_in_bicategory_axioms B.equivalence_map_def Adj.dual_equivalence
by blast
qed
show "⋀a. B.obj a ⟹ «d a : src⇩B ((E ∘ P) a) →⇩B src⇩B (I⇩C.map a)»"
unfolding emb.map_def by fastforce
show "⋀f. B.ide f ⟹ B.iso (ε⇩1 f)"
by simp
show "⋀f. B.ide f ⟹ «ε⇩1 f : I⇩C.map f ⋆⇩B d (src⇩B f) ⇒⇩B d (trg⇩B f) ⋆⇩B (E ∘ P) f»"
unfolding counit⇩1_def P_def emb.map_def
using P_def P_simps(1)
by (intro B.comp_in_homI' B.seqI B.hseqI') (auto simp add: B.iso_is_arr)
show "⋀μ. B.arr μ ⟹ ε⇩1 (B.cod μ) ⋅⇩B (I⇩C.map μ ⋆⇩B d (src⇩B μ))
= (d (trg⇩B μ) ⋆⇩B (E ∘ P) μ) ⋅⇩B ε⇩1 (B.dom μ)"
proof -
fix μ
assume μ: "B.arr μ"
let ?a = "src⇩B μ"
let ?b = "trg⇩B μ"
let ?f = "B.dom μ"
let ?g = "B.cod μ"
have "ε⇩1 ?g ⋅⇩B (I⇩C.map μ ⋆⇩B d ?a)
= (𝖺⇩B[d ?b, e ?b, ?g ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B ?g ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[?g ⋆⇩B d ?a]) ⋅⇩B
(μ ⋆⇩B d ?a)"
using μ counit⇩1_def P_def emb.map_def arr_char⇩S⇩b⇩C P⇩0_props(1) by simp
also have "... = (d ?b ⋆⇩B e ?b ⋆⇩B μ ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[d ?b, e ?b, ?f ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B ?f ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[?f ⋆⇩B d ?a]"
proof -
have "(𝖺⇩B[d ?b, e ?b, ?g ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B ?g ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[?g ⋆⇩B d ?a]) ⋅⇩B
(μ ⋆⇩B d ?a)
= 𝖺⇩B[d ?b, e ?b, ?g ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B ?g ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[?g ⋆⇩B d ?a] ⋅⇩B
(μ ⋆⇩B d ?a)"
using B.comp_assoc by simp
also have "... = 𝖺⇩B[d ?b, e ?b, ?g ⋆⇩B d ?a] ⋅⇩B
((η ?b ⋆⇩B ?g ⋆⇩B d ?a) ⋅⇩B
(?b ⋆⇩B μ ⋆⇩B d ?a)) ⋅⇩B
𝗅⇩B⇧-⇧1[?f ⋆⇩B d ?a]"
using μ B.lunit'_naturality [of "μ ⋆⇩B d ?a"] B.comp_assoc by simp
also have "... = (𝖺⇩B[d ?b, e ?b, ?g ⋆⇩B d ?a] ⋅⇩B
((d ?b ⋆⇩B e ?b) ⋆⇩B μ ⋆⇩B d ?a)) ⋅⇩B
(η ?b ⋆⇩B ?f ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[?f ⋆⇩B d ?a]"
proof -
have "(η ?b ⋆⇩B ?g ⋆⇩B d ?a) ⋅⇩B (?b ⋆⇩B μ ⋆⇩B d ?a) =
η ?b ⋅⇩B ?b ⋆⇩B ?g ⋅⇩B μ ⋆⇩B d ?a ⋅⇩B d ?a"
using μ B.iso_is_arr
B.interchange [of "η ?b" ?b "(?g ⋆⇩B d ?a)" "μ ⋆⇩B d ?a"]
B.interchange [of "?g" μ "d ?a" "d ?a"]
by simp
also have "... = (d ?b ⋆⇩B e ?b) ⋅⇩B η ?b ⋆⇩B μ ⋅⇩B ?f ⋆⇩B d ?a ⋅⇩B d ?a"
using μ B.comp_arr_dom B.comp_cod_arr apply simp
by (metis B.arrI equivalence_data_in_hom⇩B(7) equivalence_data_simps⇩B(17-18)
B.obj_trg)
also have "... = ((d ?b ⋆⇩B e ?b) ⋆⇩B μ ⋆⇩B d ?a) ⋅⇩B (η ?b ⋆⇩B ?f ⋆⇩B d ?a)"
using μ B.iso_is_arr B.interchange [of "d ?b ⋆⇩B e ?b" "η ?b" "μ ⋆⇩B d ?a" "?f ⋆⇩B d ?a"]
B.interchange [of μ "?f" "d ?a" "d ?a"]
by simp
finally have "(η ?b ⋆⇩B ?g ⋆⇩B d ?a) ⋅⇩B (?b ⋆⇩B μ ⋆⇩B d ?a)
= ((d ?b ⋆⇩B e ?b) ⋆⇩B μ ⋆⇩B d ?a) ⋅⇩B (η ?b ⋆⇩B ?f ⋆⇩B d ?a)"
by blast
thus ?thesis using B.comp_assoc by simp
qed
also have "... = (d ?b ⋆⇩B e ?b ⋆⇩B μ ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[d ?b, e ?b, ?f ⋆⇩B d ?a] ⋅⇩B
(η ?b ⋆⇩B ?f ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[?f ⋆⇩B d ?a]"
using μ B.assoc_naturality [of "d ?b" "e ?b" "μ ⋆⇩B d ?a"] B.comp_assoc by simp
finally show ?thesis by simp
qed
also have "... = (d ?b ⋆⇩B (E ∘ P) μ) ⋅⇩B ε⇩1 ?f"
using μ counit⇩1_def P_def emb.map_def arr_char⇩S⇩b⇩C P⇩0_props(1) P_simps⇩B(1)
by (simp add: P⇩0_props(6))
finally show "ε⇩1 (?g) ⋅⇩B (I⇩C.map μ ⋆⇩B d ?a) = (d ?b ⋆⇩B (E ∘ P) μ) ⋅⇩B ε⇩1 ?f"
by blast
qed
show "⋀f g. ⟦B.ide f; B.ide g; src⇩B g = trg⇩B f⟧ ⟹
(d (trg⇩B g) ⋆⇩B EoP.cmp (g, f)) ⋅⇩B
𝖺⇩B[d (trg⇩B g), (E ∘ P) g, (E ∘ P) f] ⋅⇩B
(ε⇩1 g ⋆⇩B (E ∘ P) f) ⋅⇩B
B.inv 𝖺⇩B[I⇩C.map g, d (src⇩B g), (E ∘ P) f] ⋅⇩B (I⇩C.map g ⋆⇩B ε⇩1 f) ⋅⇩B
𝖺⇩B[I⇩C.map g, I⇩C.map f, d (src⇩B f)]
= ε⇩1 (g ⋆⇩B f) ⋅⇩B (I⇩C.cmp (g, f) ⋆⇩B d (src⇩B f))"
proof -
fix f g
assume f: "B.ide f" and g: "B.ide g" and fg: "src⇩B g = trg⇩B f"
let ?a = "src⇩B f"
let ?b = "src⇩B g"
let ?c = "trg⇩B g"
have "(d ?c ⋆⇩B EoP.cmp (g, f)) ⋅⇩B
𝖺⇩B[d ?c, (E ∘ P) g, (E ∘ P) f] ⋅⇩B
(ε⇩1 g ⋆⇩B (E ∘ P) f) ⋅⇩B
B.inv 𝖺⇩B[I⇩C.map g, d ?b, (E ∘ P) f] ⋅⇩B
(I⇩C.map g ⋆⇩B ε⇩1 f) ⋅⇩B
𝖺⇩B[I⇩C.map g, I⇩C.map f, d ?a]
= (d ?c ⋆⇩B P (g ⋆⇩B f) ⋅⇩B CMP.map (g, f) ⋅⇩B (P g ⋆⇩B P f)) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
(ε⇩1 g ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B
(g ⋆⇩B ε⇩1 f) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "src⇩B (e ?c ⋆⇩B g ⋆⇩B d ?b) = trg⇩B (e ?b ⋆⇩B f ⋆⇩B d ?a)"
using f g fg src_def trg_def arr_char⇩S⇩b⇩C P_simps⇩B(1)
by (simp add: P⇩0_props(1,6) arr_char⇩S⇩b⇩C)
moreover have "B.inv 𝖺⇩B[g, d ?b, P f] = 𝖺⇩B⇧-⇧1[g, d ?b, P f]"
using f g fg by (simp add: P_def)
ultimately show ?thesis
using f g fg emb.map_def P.preserves_reflects_arr P.preserves_reflects_arr
P.preserves_reflects_arr Φ⇩P_def emb.cmp_def
EoP.cmp_def B.VV.arr_char⇩S⇩b⇩C B.VV.dom_char⇩S⇩b⇩C VV.arr_char⇩S⇩b⇩C
by simp
qed
also have "... = (d ?c ⋆⇩B CMP.map (g, f)) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
((ε⇩1 g ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[g, d ?b, P f] ⋅⇩B
(g ⋆⇩B ε⇩1 f)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "«CMP.map (g, f) : P g ⋆ P f ⇒ P (g ⋆⇩B f)»"
using f g fg VV.arr_char⇩S⇩b⇩C VV.cod_char⇩S⇩b⇩C P.FF_def Φ⇩P_def Φ⇩P_in_hom(2) by auto
hence "«CMP.map (g, f) : P g ⋆⇩B P f ⇒⇩B P (g ⋆⇩B f)»"
using in_hom_char⇩S⇩b⇩C hcomp_char by auto
hence "P (g ⋆⇩B f) ⋅⇩B CMP.map (g, f) ⋅⇩B (P g ⋆⇩B P f) = CMP.map (g, f)"
using B.comp_arr_dom B.comp_cod_arr by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c ⋆⇩B CMP.map (g, f)) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
(𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
using f g fg technical B.comp_assoc by simp
also have "... = (d ?c ⋆⇩B (e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f] ⋅⇩B
𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f] ⋅⇩B
(𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f)) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
(𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
using f g fg CMP.map_simp_ide CMP_def B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C
B.whisker_left B.whisker_left B.comp_assoc
by simp
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
((d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
(𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a])) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
using f g fg B.whisker_left P_def B.comp_assoc by simp
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
((d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
(𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a])
= 𝖺⇩B[d ?c, e ?c, g ⋆⇩B (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a]"
proof -
have "(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝖺⇩B⇧-⇧1[d ?b, e ?b, f ⋆⇩B d ?a]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B[e ?c, g, d ?b ⋆⇩B P f]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B[e ?c ⋆⇩B g, d ?b, P f]) ⋅⇩B
(d ?c ⋆⇩B 𝖺⇩B⇧-⇧1[e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B[d ?c, P g, P f] ⋅⇩B
(𝖺⇩B[d ?c, e ?c, g ⋆⇩B d ?b] ⋆⇩B P f) ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, d ?b] ⋆⇩B P f) ⋅⇩B
𝖺⇩B⇧-⇧1[(d ?c ⋆⇩B e ?c) ⋆⇩B g, d ?b, P f] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B d ?b ⋆⇩B P f) ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B 𝖺⇩B[d ?b, e ?b, f ⋆⇩B d ?a])
= ⦃(❙⟨d ?c❙⟩ ❙⋆ ❙⟨e ?c❙⟩ ❙⋆ ❙⟨g❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨d ?b❙⟩, ❙⟨e ?b❙⟩, ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙]) ❙⋅
(❙⟨d ?c❙⟩ ❙⋆ ❙𝖺❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩ ❙⋆ ❙P❙R❙J f❙]) ❙⋅
(❙⟨d ?c❙⟩ ❙⋆ ❙𝖺❙[❙⟨e ?c❙⟩ ❙⋆ ❙⟨g❙⟩, ❙⟨d ?b❙⟩, ❙P❙R❙J f❙]) ❙⋅
(❙⟨d ?c❙⟩ ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩❙] ❙⋆ ❙P❙R❙J f) ❙⋅
❙𝖺❙[❙⟨d ?c❙⟩, ❙P❙R❙J g, ❙P❙R❙J f❙] ❙⋅
(❙𝖺❙[❙⟨d ?c❙⟩, ❙⟨e ?c❙⟩, ❙⟨g❙⟩ ❙⋆ ❙⟨d ?b❙⟩❙] ❙⋆ ❙P❙R❙J f) ❙⋅
(❙𝖺❙[❙⟨d ?c❙⟩ ❙⋆ ❙⟨e ?c❙⟩, ❙⟨g❙⟩, ❙⟨d ?b❙⟩❙] ❙⋆ ❙P❙R❙J f) ❙⋅
❙𝖺⇧-⇧1❙[(❙⟨d ?c❙⟩ ❙⋆ ❙⟨e ?c❙⟩) ❙⋆ ❙⟨g❙⟩, ❙⟨d ?b❙⟩, ❙P❙R❙J f❙] ❙⋅
(((❙⟨d ?c❙⟩ ❙⋆ ❙⟨e ?c❙⟩) ❙⋆ ❙⟨g❙⟩) ❙⋆ ❙⟨d ?b❙⟩ ❙⋆ ❙P❙R❙J f) ❙⋅
(((❙⟨d ?c❙⟩ ❙⋆ ❙⟨e ?c❙⟩) ❙⋆ ❙⟨g❙⟩) ❙⋆
❙𝖺❙[❙⟨d ?b❙⟩, ❙⟨e ?b❙⟩, ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙])⦄"
using f g fg B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def B.𝔩_ide_simp
by auto
also have "... = ⦃❙𝖺❙[❙⟨d ?c❙⟩, ❙⟨e ?c❙⟩,
❙⟨g❙⟩ ❙⋆ (❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩) ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙] ❙⋅
❙𝖺❙[❙⟨d ?c❙⟩ ❙⋆ ❙⟨e ?c❙⟩, ❙⟨g❙⟩,
(❙⟨d ?b❙⟩ ❙⋆ ❙⟨e ?b❙⟩) ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙]⦄"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = 𝖺⇩B[d ?c, e ?c, g ⋆⇩B (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a]"
using f g fg B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def B.𝔩_ide_simp
by auto
finally show ?thesis by blast
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a]) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a] =
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
((d ?c ⋆⇩B e ?c) ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a)"
using f g fg
B.assoc_naturality
[of "d ?c" "e ?c" "g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
((((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "((d ?c ⋆⇩B e ?c) ⋆⇩B g ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, (d ?b ⋆⇩B e ?b) ⋆⇩B f ⋆⇩B d ?a] =
𝖺⇩B[d ?c ⋆⇩B e ?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a)"
using f g fg
B.assoc_naturality
[of "d ?c ⋆⇩B e ?c" g "B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
𝖺⇩B[d ?c ⋆⇩B e ?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a)
= (η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a"
proof -
have "(((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋆⇩B B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
((η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B η ?b ⋆⇩B f ⋆⇩B d ?a)
= ((d ?c ⋆⇩B e ?c) ⋆⇩B g) ⋅⇩B (η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B
(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)"
using f g fg
B.interchange [of "(d ?c ⋆⇩B e ?c) ⋆⇩B g" "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]"
"B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a"
"η ?b ⋆⇩B f ⋆⇩B d ?a"]
by fastforce
also have "... = (η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B
(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)"
using f g fg B.comp_cod_arr [of "(η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g]" "(d ?c ⋆⇩B e ?c) ⋆⇩B g"]
by (auto simp add: B.iso_is_arr)
also have "... = (η ?c ⋆⇩B g) ⋅⇩B 𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a"
proof -
have "(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)
= B.inv (η ?b) ⋅⇩B η ?b ⋆⇩B f ⋆⇩B d ?a"
proof -
have "B.seq (B.inv (η ?b)) (η ?b)"
using f g fg chosen_adjoint_equivalence(5) by fastforce
thus ?thesis
using f g fg B.whisker_right by simp
qed
also have "... = ?b ⋆⇩B f ⋆⇩B d ?a"
using f g fg chosen_adjoint_equivalence(5) B.comp_inv_arr' by simp
finally have "(B.inv (η ?b) ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (η ?b ⋆⇩B f ⋆⇩B d ?a)
= ?b ⋆⇩B f ⋆⇩B d ?a"
by blast
thus ?thesis
using B.comp_assoc by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝖺⇩B[d ?c ⋆⇩B e ?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
((η ?c ⋆⇩B g) ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
using f g fg B.comp_assoc
B.whisker_right [of "?b ⋆⇩B f ⋆⇩B d ?a" "η ?c ⋆⇩B g" "𝗅⇩B⇧-⇧1[g]"]
by fastforce
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
((d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a]) ⋅⇩B
(η ?c ⋆⇩B g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
using f g fg B.iso_is_arr B.comp_assoc
B.assoc_naturality [of "η ?c" g "?b ⋆⇩B f ⋆⇩B d ?a"]
by simp
also have "... = (d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
(η ?c ⋆⇩B g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "(d ?c ⋆⇩B e ?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a] =
𝖺⇩B[d ?c, e ?c, g ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
((d ?c ⋆⇩B e ?c) ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a])"
using f g fg B.iso_is_arr
B.assoc_naturality [of "d ?c" "e ?c" "g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]"]
by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ((d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B f ⋆⇩B d ?a]) ⋅⇩B
(η ?c ⋆⇩B g ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "((d ?c ⋆⇩B e ?c) ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B (η ?c ⋆⇩B g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a)
= (η ?c ⋆⇩B g ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B (?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a])"
using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr
B.interchange [of "d ?c ⋆⇩B e ?c" "η ?c"
"g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]" "g ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a"]
B.interchange [of "η ?c" ?c
"g ⋆⇩B f ⋆⇩B d ?a" "g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = 𝖺⇩B[d ?c, e ?c, (g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B
(((d ?c ⋆⇩B e ?c) ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(η ?c ⋆⇩B g ⋆⇩B f ⋆⇩B d ?a)) ⋅⇩B
(?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "(d ?c ⋆⇩B e ?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
𝖺⇩B[d ?c, e ?c, g ⋆⇩B f ⋆⇩B d ?a] =
𝖺⇩B[d ?c, e ?c, (g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B
((d ?c ⋆⇩B e ?c) ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a])"
using f g fg B.assoc_naturality [of "d ?c" "e ?c" "𝖺⇩B⇧-⇧1[g, f, d ?a]"] by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = 𝖺⇩B[d ?c, e ?c, (g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B
(η ?c ⋆⇩B (g ⋆⇩B f) ⋆⇩B d ?a) ⋅⇩B
(?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "((d ?c ⋆⇩B e ?c) ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B (η ?c ⋆⇩B g ⋆⇩B f ⋆⇩B d ?a)
= (η ?c ⋆⇩B (g ⋆⇩B f) ⋆⇩B d ?a) ⋅⇩B (?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a])"
using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr
B.interchange [of "d ?c ⋆⇩B e ?c" "η ?c" "𝖺⇩B⇧-⇧1[g, f, d ?a]" "g ⋆⇩B f ⋆⇩B d ?a"]
B.interchange [of "η ?c" ?c "(g ⋆⇩B f) ⋆⇩B d ?a" "𝖺⇩B⇧-⇧1[g, f, d ?a]"]
by auto
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = ε⇩1 (g ⋆⇩B f) ⋅⇩B (I⇩C.cmp (g, f) ⋆⇩B d ?a)"
proof -
have "ε⇩1 (g ⋆⇩B f) ⋅⇩B (I⇩C.cmp (g, f) ⋆⇩B d ?a)
= 𝖺⇩B[d ?c, e ?c, (g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B
(η ?c ⋆⇩B (g ⋆⇩B f) ⋆⇩B d ?a) ⋅⇩B
𝗅⇩B⇧-⇧1[(g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B d ?a)"
using f g fg counit⇩1_def B.comp_assoc by simp
also have "... = 𝖺⇩B[d ?c, e ?c, (g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B
(η ?c ⋆⇩B (g ⋆⇩B f) ⋆⇩B d ?a) ⋅⇩B
(?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "𝗅⇩B⇧-⇧1[(g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B d ?a)
= (?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
proof -
have "𝗅⇩B⇧-⇧1[(g ⋆⇩B f) ⋆⇩B d ?a] ⋅⇩B ((g ⋆⇩B f) ⋆⇩B d ?a) =
⦃❙𝗅⇧-⇧1❙[(❙⟨g❙⟩ ❙⋆ ❙⟨f❙⟩) ❙⋆ ❙⟨d ?a❙⟩❙] ❙⋅ ((❙⟨g❙⟩ ❙⋆ ❙⟨f❙⟩) ❙⋆ ❙⟨d ?a❙⟩)⦄"
using f g fg B.𝔩_ide_simp by auto
also have "... = ⦃(❙⟨?c❙⟩⇩0 ❙⋆ ❙𝖺⇧-⇧1❙[❙⟨g❙⟩, ❙⟨f❙⟩, ❙⟨d ?a❙⟩❙]) ❙⋅
(❙⟨?c❙⟩⇩0 ❙⋆ ❙⟨g❙⟩ ❙⋆ ❙𝗅❙[❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙]) ❙⋅
❙𝖺❙[❙⟨?c❙⟩⇩0, ❙⟨g❙⟩, ❙⟨?b❙⟩⇩0 ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩❙] ❙⋅
(❙𝗅⇧-⇧1❙[❙⟨g❙⟩❙] ❙⋆ ❙⟨?b❙⟩⇩0 ❙⋆ ❙⟨f❙⟩ ❙⋆ ❙⟨d ?a❙⟩) ❙⋅
(❙⟨g❙⟩ ❙⋆ ❙𝖺❙[❙⟨?b❙⟩⇩0, ❙⟨f❙⟩, ❙⟨d ?a❙⟩❙] ❙⋅ (❙𝗅⇧-⇧1❙[❙⟨f❙⟩❙] ❙⋆ ❙⟨d ?a❙⟩)) ❙⋅
❙𝖺❙[❙⟨g❙⟩, ❙⟨f❙⟩, ❙⟨d ?a❙⟩❙]⦄"
using f g fg by (intro EV.eval_eqI, auto)
also have "... = (?c ⋆⇩B 𝖺⇩B⇧-⇧1[g, f, d ?a]) ⋅⇩B
(?c ⋆⇩B g ⋆⇩B 𝗅⇩B[f ⋆⇩B d ?a]) ⋅⇩B
𝖺⇩B[?c, g, ?b ⋆⇩B f ⋆⇩B d ?a] ⋅⇩B
(𝗅⇩B⇧-⇧1[g] ⋆⇩B ?b ⋆⇩B f ⋆⇩B d ?a) ⋅⇩B
(g ⋆⇩B 𝖺⇩B[?b, f, d ?a] ⋅⇩B (𝗅⇩B⇧-⇧1[f] ⋆⇩B d ?a)) ⋅⇩B
𝖺⇩B[g, f, d ?a]"
using f g fg B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def B.𝔩_ide_simp
by auto
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
finally show "(d ?c ⋆⇩B EoP.cmp (g, f)) ⋅⇩B
𝖺⇩B[d ?c, (E ∘ P) g, (E ∘ P) f] ⋅⇩B
(ε⇩1 g ⋆⇩B (E ∘ P) f) ⋅⇩B
B.inv 𝖺⇩B[I⇩C.map g, d ?b, (E ∘ P) f] ⋅⇩B (I⇩C.map g ⋆⇩B ε⇩1 f) ⋅⇩B
𝖺⇩B[I⇩C.map g, I⇩C.map f, d ?a]
= ε⇩1 (g ⋆⇩B f) ⋅⇩B (I⇩C.cmp (g, f) ⋆⇩B d ?a)"
by blast
qed
show "⋀a. B.obj a ⟹
(d a ⋆⇩B EoP.unit a) ⋅⇩B 𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a] = ε⇩1 a ⋅⇩B (I⇩C.unit a ⋆⇩B d a)"
proof -
fix a
assume a: "B.obj a"
interpret adjoint_equivalence_in_bicategory V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B ‹e a› ‹d a› ‹η a› ‹ε a›
using a chosen_adjoint_equivalence by simp
have 0: "src⇩B a = a ∧ trg⇩B a = a"
using a by auto
have 1: "obj (P⇩0 a)"
using a P⇩0_props(1) by simp
have "(d a ⋆⇩B EoP.unit a) ⋅⇩B 𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a]
= (d a ⋆⇩B ((e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B src⇩B (a ⋆⇩B d (src⇩B a))) ⋅⇩B
𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a]"
proof -
have "B.hseq (e (trg⇩B a)) (a ⋆⇩B d (src⇩B a))"
using a by (elim B.objE, intro B.hseqI') auto
moreover have "arr (src⇩B (d (src⇩B a)))"
using a arr_char⇩S⇩b⇩C P⇩0_props(1) obj_char B.obj_simps(1) by auto
moreover have "arr (src⇩B (a ⋆⇩B d (src⇩B a)))"
using a arr_char⇩S⇩b⇩C P⇩0_props(1) obj_char calculation(1) by fastforce
moreover have "src⇩B (a ⋆⇩B d (src⇩B a)) ∈ Obj ∧ trg⇩B (e (trg⇩B a)) ∈ Obj"
using a B.obj_simps P⇩0_props obj_char arr_char⇩S⇩b⇩C by simp
moreover have "P⇩0 a ⋅⇩B P⇩0 a ∈ Obj"
using 1 arr_char⇩S⇩b⇩C P⇩0_props(1) obj_char
by (metis (no_types, lifting) B.cod_trg B.obj_def' B.trg.as_nat_trans.naturality2)
moreover have "emb.unit (P⇩0 (src⇩B a)) = P⇩0 (src⇩B a)"
using a 0 1 emb.unit_char' P.map⇩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⇩0_def P_def
src_def arr_char⇩S⇩b⇩C obj_char
by simp
qed
also have "... = (d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a) ⋅⇩B P⇩0 a) ⋅⇩B
𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a]"
using a 1 B.comp_assoc B.obj_simps by auto
also have "... = (d a ⋆⇩B (e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B B.inv (ε a)) ⋅⇩B 𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a]"
using a B.comp_arr_dom by simp
also have "... = (d a ⋆⇩B e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B (d a ⋆⇩B B.inv (ε a)) ⋅⇩B 𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a]"
using a B.whisker_left B.comp_assoc by simp
also have "... = ((d a ⋆⇩B e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝖺⇩B[d a, e a, d a]) ⋅⇩B (η a ⋆⇩B d a)"
using a triangle_right B.comp_assoc
B.invert_side_of_triangle(1) [of "𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a]" "d a ⋆⇩B ε a"]
by simp
also have "... = 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B ((d a ⋆⇩B e a) ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B (η a ⋆⇩B d a)"
proof -
have "(d a ⋆⇩B e a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]) ⋅⇩B 𝖺⇩B[d a, e a, d a] =
𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B ((d a ⋆⇩B e a) ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using a B.assoc_naturality [of "d a" "e a" "𝗅⇩B⇧-⇧1[d a]"] by simp
thus ?thesis
using B.comp_assoc by simp
qed
also have "... = 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using a B.interchange [of "d a ⋆⇩B e a" "η a" "𝗅⇩B⇧-⇧1[d a]" "d a"]
B.comp_arr_dom B.comp_cod_arr
by simp
also have "... = 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B[a, a, d a] ⋅⇩B
(𝗅⇩B⇧-⇧1[a] ⋆⇩B d a)"
proof -
have "𝖺⇩B[a, a, d a] ⋅⇩B (𝗅⇩B⇧-⇧1[a] ⋆⇩B d a) = a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]"
proof -
have "𝖺⇩B[a, a, d a] ⋅⇩B (𝗅⇩B⇧-⇧1[a] ⋆⇩B d a) = ⦃❙𝖺❙[❙⟨a❙⟩⇩0, ❙⟨a❙⟩⇩0, ❙⟨d a❙⟩❙] ❙⋅ (❙𝗅⇧-⇧1❙[❙⟨a❙⟩⇩0❙] ❙⋆ ❙⟨d a❙⟩)⦄"
using a ide_char B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def B.𝔩_ide_simp
by auto
also have "... = ⦃❙⟨a❙⟩⇩0 ❙⋆ ❙𝗅⇧-⇧1❙[❙⟨d a❙⟩❙]⦄"
using a ide_char⇩S⇩b⇩C by (intro EV.eval_eqI, auto)
also have "... = a ⋆⇩B 𝗅⇩B⇧-⇧1[d a]"
using a ide_char B.α_def B.α'.map_ide_simp B.VVV.ide_char⇩S⇩b⇩C B.VVV.arr_char⇩S⇩b⇩C
B.VV.ide_char⇩S⇩b⇩C B.VV.arr_char⇩S⇩b⇩C P_def B.𝔩_ide_simp
by auto
finally show ?thesis by blast
qed
hence "𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B[a, a, d a] ⋅⇩B
(𝗅⇩B⇧-⇧1[a] ⋆⇩B d a)
= 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B (a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
by simp
also have "... = 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B 𝗅⇩B⇧-⇧1[d a])"
using a B.interchange [of "η a" a "a ⋆⇩B d a" "𝗅⇩B⇧-⇧1[d a]"] B.comp_arr_dom B.comp_cod_arr
by simp
finally show ?thesis by simp
qed
also have "... = ε⇩1 a ⋅⇩B (I⇩C.unit a ⋆⇩B d a)"
proof -
have "ε⇩1 a ⋅⇩B (I⇩C.unit a ⋆⇩B d a) =
𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝗅⇩B⇧-⇧1[a ⋆⇩B d a] ⋅⇩B (a ⋆⇩B d a)"
using a 0 counit⇩1_def I⇩C.unit_char' B.comp_assoc by simp
also have "... = 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝗅⇩B⇧-⇧1[a ⋆⇩B d a]"
proof -
have "B.arr 𝗅⇩B⇧-⇧1[a ⋆⇩B d a]"
using a
by (metis B.ide_hcomp B.lunit'_simps(1) B.objE equivalence_data_simps⇩B(8) ide_right)
moreover have "B.dom 𝗅⇩B⇧-⇧1[a ⋆⇩B d a] = a ⋆⇩B d a"
using a
by (metis B.ide_hcomp B.lunit'_simps(4) B.objE antipar(1) equivalence_data_simps⇩B(5)
ide_right)
ultimately show ?thesis
using a B.comp_arr_dom [of "𝗅⇩B⇧-⇧1[a ⋆⇩B d a]" "a ⋆⇩B d a"] by simp
qed
also have "... = 𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B[a, a, d a] ⋅⇩B
(𝗅⇩B⇧-⇧1[a] ⋆⇩B d a)"
using a B.lunit_hcomp(4) [of a "d a"] by auto
finally have "ε⇩1 a ⋅⇩B (I⇩C.unit a ⋆⇩B d a) =
𝖺⇩B[d a, e a, a ⋆⇩B d a] ⋅⇩B (η a ⋆⇩B a ⋆⇩B d a) ⋅⇩B 𝖺⇩B[a, a, d a] ⋅⇩B
(𝗅⇩B⇧-⇧1[a] ⋆⇩B d a)"
by simp
thus ?thesis by simp
qed
finally show "(d a ⋆⇩B EoP.unit a) ⋅⇩B 𝗋⇩B⇧-⇧1[d a] ⋅⇩B 𝗅⇩B[d a] = ε⇩1 a ⋅⇩B (I⇩C.unit a ⋆⇩B d a)"
by blast
qed
qed
interpretation equivalence_of_bicategories V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B comp hcomp 𝖺 𝗂⇩B src trg
E Φ⇩E P Φ⇩P unit⇩0 unit⇩1 counit⇩0 counit⇩1
..
lemma induces_equivalence:
shows "equivalence_of_bicategories V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B comp hcomp 𝖺 𝗂⇩B src trg
E Φ⇩E P Φ⇩P unit⇩0 unit⇩1 counit⇩0 counit⇩1"
..
end
subsection "Equivalence Pseudofunctors, Bijective on Objects"
text ‹
Here we carry out the extension of an equivalence pseudofunctor ‹F› to an equivalence
of bicategories in the special case that the object map of ‹F› is bijective.
The bijectivity assumption simplifies the construction of the unit and counit of the
equivalence (the components at objects are in fact identities), as well as the proofs
of the associated coherence conditions.
›
locale equivalence_pseudofunctor_bij_on_obj =
equivalence_pseudofunctor +
assumes bij_on_obj: "bij_betw map⇩0 (Collect C.obj) (Collect D.obj)"
begin
abbreviation F⇩0
where "F⇩0 ≡ map⇩0"
definition G⇩0
where "G⇩0 = inv_into (Collect C.obj) F⇩0"
lemma G⇩0_props:
shows "D.obj b ⟹ C.obj (G⇩0 b) ∧ F⇩0 (G⇩0 b) = b"
and "C.obj a ⟹ D.obj (F⇩0 a) ∧ G⇩0 (F⇩0 a) = a"
proof -
have surj: "F⇩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⇩0 b) ∧ F⇩0 (G⇩0 b) = b"
proof
show "C.obj (G⇩0 b)"
unfolding G⇩0_def
using b by (metis inv_into_into mem_Collect_eq surj)
show "F⇩0 (G⇩0 b) = b"
unfolding G⇩0_def
using b by (simp add: f_inv_into_f surj)
qed
next
have bij: "bij_betw G⇩0 (Collect D.obj) (Collect C.obj)"
using bij_on_obj G⇩0_def bij_betw_inv_into by auto
have surj: "G⇩0 ` (Collect D.obj) = Collect C.obj"
using bij_on_obj G⇩0_def
by (metis bij_betw_def bij_betw_inv_into)
assume a: "C.obj a"
show "D.obj (F⇩0 a) ∧ G⇩0 (F⇩0 a) = a"
using a bij surj G⇩0_def bij_betw_imp_inj_on bij_on_obj by force
qed
text ‹
We extend ‹G⇩0› to all arrows of ‹D› using chosen adjoint equivalences, which extend ‹F›,
between ‹hom⇩C (a, a')› and ‹hom⇩D (F a, F a')›. The use of adjoint equivalences restricts
choices that prevent us from validating the necessary coherence conditions.
›
definition G
where "G ν = (if D.arr ν then
equivalence_pseudofunctor_at_hom.G⇩1 V⇩C src⇩C trg⇩C V⇩D src⇩D trg⇩D F
(G⇩0 (src⇩D ν)) (G⇩0 (trg⇩D ν)) ν
else C.null)"
lemma G_in_hom [intro]:
assumes "D.arr ν"
shows "«G ν : G⇩0 (src⇩D ν) →⇩C G⇩0 (trg⇩D ν)»"
and "«G ν : G (D.dom ν) ⇒⇩C G (D.cod ν)»"
proof -
have 1: "src⇩D ν = F⇩0 (G⇩0 (src⇩D ν)) ∧ trg⇩D ν = F⇩0 (G⇩0 (trg⇩D ν))"
using assms G⇩0_props by simp
interpret hom⇩C: subcategory V⇩C ‹λμ. «μ : G⇩0 (src⇩D ν) →⇩C G⇩0 (trg⇩D ν)»›
using assms C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 (src⇩D ν)) →⇩D F⇩0 (G⇩0 (trg⇩D ν))»›
using assms 1 D.in_hhom_def D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D ν)› ‹G⇩0 (trg⇩D ν)›
using assms G⇩0_props(1) by unfold_locales auto
have 2: "hom⇩D.arr ν"
using assms 1 hom⇩D.arr_char⇩S⇩b⇩C by simp
show "«G ν : G (D.dom ν) ⇒⇩C G (D.cod ν)»"
unfolding G_def
using assms 2 hom⇩C.arr_char⇩S⇩b⇩C hom⇩C.dom_char⇩S⇩b⇩C hom⇩C.cod_char⇩S⇩b⇩C hom⇩D.dom_char⇩S⇩b⇩C hom⇩D.cod_char⇩S⇩b⇩C
Faa'.ηε.F.preserves_arr Faa'.ηε.F.preserves_dom Faa'.ηε.F.preserves_cod
by (intro C.in_homI) auto
thus "C.in_hhom (G ν) (G⇩0 (src⇩D ν)) (G⇩0 (trg⇩D ν))"
proof -
have "hom⇩C.arr (Faa'.G⇩1 ν)"
using 2 by simp
thus ?thesis
using G_def assms hom⇩C.arrE by presburger
qed
qed
lemma G_simps [simp]:
assumes "D.arr ν"
shows "C.arr (G ν)"
and "src⇩C (G ν) = G⇩0 (src⇩D ν)" and "trg⇩C (G ν) = G⇩0 (trg⇩D ν)"
and "C.dom (G ν) = G (D.dom ν)" and "C.cod (G ν) = G (D.cod ν)"
using assms G_in_hom by auto
interpretation G: "functor" V⇩D V⇩C G
proof
show "⋀f. ¬ D.arr f ⟹ G f = C.null"
unfolding G_def by simp
fix ν
assume ν: "D.arr ν"
show "C.arr (G ν)"
using ν by simp
show "C.dom (G ν) = G (D.dom ν)"
using ν by simp
show "C.cod (G ν) = G (D.cod ν)"
using ν by simp
next
fix μ ν
assume μν: "D.seq μ ν"
have 1: "D.arr μ ∧ D.arr ν ∧ src⇩D μ = src⇩D ν ∧ trg⇩D μ = trg⇩D ν ∧
src⇩D μ = F⇩0 (G⇩0 (src⇩D μ)) ∧ trg⇩D μ = F⇩0 (G⇩0 (trg⇩D μ)) ∧
src⇩D ν = F⇩0 (G⇩0 (src⇩D ν)) ∧ trg⇩D ν = F⇩0 (G⇩0 (trg⇩D ν))"
using μν G⇩0_props D.vseq_implies_hpar by auto
interpret hom⇩C: subcategory V⇩C ‹λμ. «μ : G⇩0 (src⇩D ν) →⇩C G⇩0 (trg⇩D ν)»›
using 1 C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 (src⇩D ν)) →⇩D F⇩0 (G⇩0 (trg⇩D ν))»›
using 1 D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret hom⇩D': subcategory V⇩D
‹λν'. D.arr ν' ∧ src⇩D ν' = src⇩D ν ∧ trg⇩D ν' = trg⇩D ν›
using 1 D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D ν)› ‹G⇩0 (trg⇩D ν)›
using μν 1 G⇩0_props(1) by unfold_locales auto
have 2: "hom⇩D.seq μ ν"
using μν 1 hom⇩D.arr_char⇩S⇩b⇩C hom⇩D.comp_def by fastforce
have "G (hom⇩D.comp μ ν) = hom⇩C.comp (G μ) (G ν)"
unfolding G_def
using 1 2 G⇩0_props Faa'.G⇩1_props hom⇩D.arr_char⇩S⇩b⇩C
Faa'.ηε.F.as_nat_trans.preserves_comp_2
by auto
thus "G (μ ⋅⇩D ν) = G μ ⋅⇩C G ν"
using μν 1 2 G_def hom⇩C.comp_simp hom⇩D.comp_simp D.src_vcomp D.trg_vcomp
Faa'.ηε.F.preserves_arr
by metis
qed
lemma functor_G:
shows "functor V⇩D V⇩C G"
..
interpretation G: faithful_functor V⇩D V⇩C G
proof
fix f f'
assume par: "D.par f f'"
assume eq: "G f = G f'"
have 1: "src⇩D f = src⇩D f' ∧ trg⇩D f = trg⇩D f'"
using par by (metis D.src_dom D.trg_dom)
interpret hom⇩C: subcategory V⇩C ‹λμ. «μ : G⇩0 (src⇩D f) →⇩C G⇩0 (trg⇩D f)»›
using par C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 (src⇩D f)) →⇩D F⇩0 (G⇩0 (trg⇩D f))»›
using par D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret F: equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D f)› ‹G⇩0 (trg⇩D f)›
using par G⇩0_props(1) by unfold_locales auto
interpret F': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D f')› ‹G⇩0 (trg⇩D f')›
using par G⇩0_props(1) by unfold_locales auto
have 2: "hom⇩D.par f f'"
using par 1 hom⇩D.arr_char⇩S⇩b⇩C hom⇩D.dom_char⇩S⇩b⇩C hom⇩D.cod_char⇩S⇩b⇩C G⇩0_props by simp
have "F.G⇩1 f = F.G⇩1 f'"
using par eq G_def G_simps(2-3) by metis
thus "f = f'"
using F.ηε.F_is_faithful
by (simp add: 2 faithful_functor_axioms_def faithful_functor_def)
qed
interpretation FG: composite_functor V⇩D V⇩C V⇩D G F ..
interpretation FG: faithful_functor V⇩D V⇩D "F o G"
using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
by blast
interpretation GF: composite_functor V⇩C V⇩D V⇩C F G ..
interpretation GF: faithful_functor V⇩C V⇩C "G o F"
using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
by blast
interpretation G: weak_arrow_of_homs V⇩D src⇩D trg⇩D V⇩C src⇩C trg⇩C G
proof
have *: "⋀b. D.obj b ⟹ C.isomorphic (G b) (src⇩C (G b))"
proof -
fix b
assume b: "D.obj b"
interpret hom⇩C: subcategory V⇩C ‹λμ. «μ : G⇩0 b →⇩C G⇩0 b»›
using b C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 b) →⇩D F⇩0 (G⇩0 b)»›
using b D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 b› ‹G⇩0 b›
using b G⇩0_props(1) by unfold_locales auto
have 1: "C.in_hhom (G⇩0 b) (G⇩0 b) (G⇩0 b)"
using b G⇩0_props by force
text ‹
Using the unit constraints of ‹F› and the fact that ‹F⇩0 (G⇩0 b) = b›,
we obtain an isomorphism ‹«?ψ : b ⇒⇩D F (G⇩0 b)»›, which is also
an isomorphism in ‹hom⇩D›.
›
let ?ψ = "unit (G⇩0 b)"
have ψ: "«?ψ : b ⇒⇩D F (G⇩0 b)» ∧ D.iso ?ψ"
using b G⇩0_props unit_char by auto
have ψ_in_hhom: "D.in_hhom ?ψ b b"
using b ψ 1 D.src_dom D.trg_dom by fastforce
have 2: "hom⇩D.arr ?ψ ∧ hom⇩D.arr (D.inv ?ψ)"
by (metis D.in_hhomE D.inv_in_hhom G⇩0_props(1) ψ ψ_in_hhom hom⇩D.arr_char⇩S⇩b⇩C)
have ψ': "hom⇩D.in_hom ?ψ b (Faa'.F⇩1 (G⇩0 b))"
proof
show "hom⇩D.arr ?ψ"
using 2 by simp
show "hom⇩D.dom ?ψ = b"
using 2 ψ hom⇩D.dom_char⇩S⇩b⇩C by auto
show "hom⇩D.cod ?ψ = Faa'.F⇩1 (G⇩0 b)"
proof -
have "«?ψ : b ⇒⇩D Faa'.F⇩1 (G⇩0 b)» ∧ D.iso ?ψ"
unfolding Faa'.F⇩1_def using b ψ 1 G⇩0_props by auto
thus ?thesis
using 2 ψ hom⇩D.cod_char⇩S⇩b⇩C by auto
qed
qed
text ‹
Transposing ‹?ψ› via the adjunction ‹ηε› yields an isomorphism from
‹G b› to ‹G⇩0 b› in ‹hom⇩C›, hence in ‹C›.
›
have **: "hom⇩C.isomorphic (G b) (G⇩0 b)"
proof -
have "hom⇩C.in_hom (Faa'.ηε.ψ (G⇩0 b) ?ψ) (Faa'.G⇩1 b) (G⇩0 b)"
by (metis "1" C.ide_src C.in_hhom_def Faa'.ηε.ψ_in_hom ψ' hom⇩C.ideI⇩S⇩b⇩C)
moreover have "hom⇩C.iso (Faa'.ηε.ψ (G⇩0 b) ?ψ)"
proof (unfold Faa'.ηε.ψ_def)
have "hom⇩C.iso_in_hom (hom⇩C.comp (Faa'.ε (G⇩0 b))
(Faa'.G⇩1 ?ψ)) (Faa'.G⇩1 b) (G⇩0 b)"
proof (intro hom⇩C.comp_iso_in_hom)
show "hom⇩C.iso_in_hom (Faa'.G⇩1 (unit (G⇩0 b)))
(Faa'.G⇩1 b) (Faa'.ηε.FG.map (G⇩0 b))"
using ψ ψ' 1 2 hom⇩D.iso_char⇩S⇩b⇩C Faa'.ηε.F.preserves_iso hom⇩D.iso_char⇩S⇩b⇩C
Faa'.ηε.F.preserves_iso hom⇩C.in_hom_char⇩S⇩b⇩C hom⇩C.arr_char⇩S⇩b⇩C
by auto
show "hom⇩C.iso_in_hom (Faa'.ε (G⇩0 b)) (Faa'.ηε.FG.map (G⇩0 b)) (G⇩0 b)"
using 1 C.ide_src C.ide_trg C.in_hhom_def Faa'.ηε.ε.components_are_iso
Faa'.ηε.ε.preserves_hom hom⇩C.arrI hom⇩C.ideI⇩S⇩b⇩C hom⇩C.ide_in_hom
hom⇩C.map_simp hom⇩C.iso_in_hom_def
by metis
qed
thus "hom⇩C.iso (hom⇩C.comp (Faa'.ε (G⇩0 b)) (Faa'.G⇩1 (unit (G⇩0 b))))" by auto
qed
ultimately show ?thesis
unfolding G_def
using b hom⇩C.isomorphic_def D.obj_def D.obj_simps(3) by auto
qed
hence "hom⇩C.isomorphic (G b) (src⇩C (G b))"
using b G_in_hom(1) by auto
moreover have "⋀f g. hom⇩C.isomorphic f g ⟹ C.isomorphic f g"
using hom⇩C.iso_char⇩S⇩b⇩C hom⇩C.isomorphic_def C.isomorphic_def hom⇩C.in_hom_char⇩S⇩b⇩C hom⇩C.arr_char⇩S⇩b⇩C
by auto
ultimately show "C.isomorphic (G b) (src⇩C (G b))"
by simp
qed
fix ν
assume ν: "D.arr ν"
show "C.isomorphic (G (src⇩D ν)) (src⇩C (G ν))"
using ν * by force
show "C.isomorphic (G (trg⇩D ν)) (trg⇩C (G ν))"
using ν * by force
qed
lemma weak_arrow_of_homs_G:
shows "weak_arrow_of_homs V⇩D src⇩D trg⇩D V⇩C src⇩C trg⇩C G"
..
sublocale H⇩DoGG: composite_functor D.VV.comp C.VV.comp V⇩C
G.FF ‹λμν. fst μν ⋆⇩C snd μν›
..
sublocale GoH⇩D: composite_functor D.VV.comp V⇩D V⇩C ‹λμν. fst μν ⋆⇩D snd μν› G
..
text ‹
To get the unit ‹η› of the equivalence of bicategories, we piece together the
units from the local equivalences. The components at objects will in fact be identities.
›
definition η
where "η ν = (if D.arr ν then
equivalence_pseudofunctor_at_hom.η V⇩C src⇩C trg⇩C V⇩D src⇩D trg⇩D F
(G⇩0 (src⇩D ν)) (G⇩0 (trg⇩D ν)) ν
else D.null)"
lemma η_in_hom:
assumes "D.arr ν"
shows [intro]: "«η ν : src⇩D ν →⇩D trg⇩D ν»"
and [intro]: "«η ν : D.dom ν ⇒⇩D F (G (D.cod ν))»"
and "D.ide ν ⟹ D.iso (η ν)"
proof -
interpret hom⇩C: subcategory V⇩C ‹λμ. «μ : G⇩0 (src⇩D ν) →⇩C G⇩0 (trg⇩D ν)»›
using assms C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 (src⇩D ν)) →⇩D F⇩0 (G⇩0 (trg⇩D ν))»›
using assms D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D ν)› ‹G⇩0 (trg⇩D ν)›
using assms G⇩0_props(1) by unfold_locales auto
have 1: "«η ν : D.dom ν ⇒⇩D F (G (D.cod ν))» ∧ (D.ide ν ⟶ D.iso (η ν))"
proof -
have "src⇩D ν = F⇩0 (G⇩0 (src⇩D ν)) ∧ trg⇩D ν = F⇩0 (G⇩0 (trg⇩D ν))"
using assms G⇩0_props by simp
hence "hom⇩D.arr ν"
using assms hom⇩D.arr_char⇩S⇩b⇩C hom⇩D.ide_char by simp
hence "hom⇩D.in_hom (Faa'.η ν) (D.dom ν) (F (G (D.cod ν))) ∧
(D.ide ν ⟶ hom⇩D.iso (Faa'.η ν))"
using assms Faa'.ηε.η.preserves_hom [of ν "D.dom ν" "D.cod ν"]
hom⇩D.map_def G_def Faa'.F⇩1_def
apply simp
by (simp add: D.arr_iff_in_hom hom⇩D.arr_char⇩S⇩b⇩C hom⇩D.cod_closed hom⇩D.dom_closed
hom⇩D.ide_char⇩S⇩b⇩C hom⇩D.in_hom_char⇩S⇩b⇩C)
thus ?thesis
unfolding η_def
using hom⇩D.in_hom_char⇩S⇩b⇩C hom⇩D.iso_char⇩S⇩b⇩C by auto
qed
show "«η ν : D.dom ν ⇒⇩D F (G (D.cod ν))»"
using 1 by simp
thus "D.in_hhom (η ν) (src⇩D ν) (trg⇩D ν)"
using assms D.src_dom D.trg_dom
by (metis D.arrI D.in_hhom_def D.vconn_implies_hpar(1-2))
show "D.ide ν ⟹ D.iso (η ν)"
using 1 by simp
qed
lemma η_simps [simp]:
assumes "D.arr ν"
shows "D.arr (η ν)"
and "src⇩D (η ν) = src⇩D ν" and "trg⇩D (η ν) = trg⇩D ν"
and "D.dom (η ν) = D.dom ν" and "D.cod (η ν) = F (G (D.cod ν))"
and "D.ide ν ⟹ D.iso (η ν)"
using assms η_in_hom by auto
lemma η_naturality:
assumes "D.arr ν"
shows "η (D.cod ν) ⋅⇩D ν = η ν" and "F (G ν) ⋅⇩D η (D.dom ν) = η ν"
and "ν ⋅⇩D D.inv (η (D.dom ν)) = D.inv (η (D.cod ν)) ⋅⇩D F (G ν)"
proof -
interpret hom⇩C: subcategory V⇩C ‹λμ'. «μ' : G⇩0 (src⇩D ν) →⇩C G⇩0 (trg⇩D ν)»›
using assms C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 (src⇩D ν)) →⇩D F⇩0 (G⇩0 (trg⇩D ν))»›
using D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D ν)› ‹G⇩0 (trg⇩D ν)›
using assms G⇩0_props by unfold_locales auto
show 1: "η (D.cod ν) ⋅⇩D ν = η ν"
unfolding η_def
using assms Faa'.ηε.η.naturality2 hom⇩D.comp_char G_def hom⇩D.cod_simp
G_in_hom(1) hom⇩C.arr_char⇩S⇩b⇩C hom⇩D.arr_char⇩S⇩b⇩C hom⇩D.cod_closed
apply simp
by (metis (no_types, lifting) D.ext Faa'.ηε.F.preserves_reflects_arr
Faa'.ηε.η.preserves_reflects_arr)
show 2: "F (G ν) ⋅⇩D η (D.dom ν) = η ν"
unfolding η_def
using assms D.src_dom D.src_cod D.trg_dom D.trg_cod Faa'.ηε.η.naturality1
hom⇩D.comp_char G_def Faa'.F⇩1_def hom⇩D.dom_simp hom⇩D.cod_simp
apply simp
by (metis (no_types, lifting) D.not_arr_null Faa'.ηε.η.extensionality
η_def η_simps(1) hom⇩D.null_char)
show "ν ⋅⇩D D.inv (η (D.dom ν)) = D.inv (η (D.cod ν)) ⋅⇩D F (G ν)"
using assms 1 2
by (simp add: D.invert_opposite_sides_of_square)
qed
text ‹
The fact that ‹G⇩0› is chosen to be right-inverse to ‹F› implies that ‹η› is an
ordinary natural isomorphism (with respect to vertical composition) from ‹I⇩D› to ‹FG›.
›
interpretation η: natural_isomorphism V⇩D V⇩D D.map FG.map η
using η_simps η_naturality η_def
by unfold_locales auto
text ‹
In view of the bijectivity assumption, we can obtain the counit ‹ε› the same way.
›
definition ε
where "ε μ = (if C.arr μ then
equivalence_pseudofunctor_at_hom.ε V⇩C src⇩C trg⇩C V⇩D src⇩D trg⇩D F
(src⇩C μ) (trg⇩C μ) μ
else C.null)"
lemma ε_in_hom:
assumes "C.arr μ"
shows [intro]: "«ε μ : src⇩C μ →⇩C trg⇩C μ»"
and [intro]: "«ε μ : G (F (C.dom μ)) ⇒⇩C C.cod μ»"
and "C.ide μ ⟹ C.iso (ε μ)"
proof -
interpret hom⇩C: subcategory V⇩C ‹λμ'. «μ' : src⇩C μ →⇩C trg⇩C μ»›
using C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (src⇩C μ) →⇩D F⇩0 (trg⇩C μ)»›
using assms D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (F⇩0 (src⇩C μ))› ‹G⇩0 (F⇩0 (trg⇩C μ))›
using assms G⇩0_props
by unfold_locales simp_all
have 1: "«ε μ : G (F (C.dom μ)) ⇒⇩C C.cod μ» ∧ (C.ide μ ⟶ C.iso (ε μ))"
proof -
have "hom⇩C.arr μ"
using assms hom⇩C.arr_char⇩S⇩b⇩C by simp
hence "hom⇩C.in_hom μ (C.dom μ) (C.cod μ)"
using assms hom⇩C.in_hom_char⇩S⇩b⇩C hom⇩C.dom_char⇩S⇩b⇩C hom⇩C.cod_char⇩S⇩b⇩C by auto
hence "hom⇩C.in_hom (Faa'.ε μ) (G (F (C.dom μ))) (C.cod μ) ∧
(C.ide μ ⟶ hom⇩C.iso (Faa'.ε μ))"
using assms Faa'.ηε.ε.preserves_hom [of μ "C.dom μ" "C.cod μ"]
Faa'.ηε.ε.components_are_iso hom⇩C.map_def G_def Faa'.F⇩1_def G⇩0_props
by (simp add: hom⇩C.ideI⇩S⇩b⇩C)
thus ?thesis
unfolding ε_def
using assms hom⇩C.in_hom_char⇩S⇩b⇩C hom⇩C.iso_char⇩S⇩b⇩C G⇩0_props(2) by simp
qed
show "«ε μ : G (F (C.dom μ)) ⇒⇩C C.cod μ»"
using 1 by simp
thus "C.in_hhom (ε μ) (src⇩C μ) (trg⇩C μ)"
using assms C.src_cod C.trg_cod
by (metis C.arrI C.in_hhom_def C.vconn_implies_hpar(1-4))
show "C.ide μ ⟹ C.iso (ε μ)"
using 1 by simp
qed
lemma ε_simps [simp]:
assumes "C.arr μ"
shows "C.arr (ε μ)"
and "src⇩C (ε μ) = src⇩C μ" and "trg⇩C (ε μ) = trg⇩C μ"
and "C.dom (ε μ) = G (F (C.dom μ))" and "C.cod (ε μ) = C.cod μ"
and "C.ide μ ⟹ C.iso (ε μ)"
using assms ε_in_hom by auto
lemma ε_naturality:
assumes "C.arr μ"
shows "ε (C.cod μ) ⋅⇩C G (F μ) = ε μ" and "μ ⋅⇩C ε (C.dom μ) = ε μ"
and "G (F μ) ⋅⇩C C.inv (ε (C.dom μ)) = C.inv (ε (C.cod μ)) ⋅⇩C μ"
proof -
interpret hom⇩C: subcategory V⇩C ‹λμ'. «μ' : src⇩C μ →⇩C trg⇩C μ»›
using assms C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν. «ν : F⇩0 (src⇩C μ) →⇩D F⇩0 (trg⇩C μ)»›
using assms D.hhom_is_subcategory
by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹src⇩C μ› ‹trg⇩C μ›
using assms by unfold_locales simp_all
show 1: "ε (C.cod μ) ⋅⇩C G (F μ) = ε μ"
unfolding ε_def
using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.ηε.ε.naturality2
hom⇩C.comp_char G_def Faa'.F⇩1_def G⇩0_props hom⇩C.dom_char⇩S⇩b⇩C hom⇩C.cod_char⇩S⇩b⇩C
apply simp
by (metis (no_types, lifting) C.in_hhomI Faa'.ηε.ε.preserves_reflects_arr hom⇩C.arr_char⇩S⇩b⇩C
hom⇩C.not_arr_null hom⇩C.null_char)
show 2: "μ ⋅⇩C ε (C.dom μ) = ε μ"
unfolding ε_def
using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.ηε.ε.naturality1
hom⇩C.comp_char G_def Faa'.F⇩1_def G⇩0_props hom⇩C.dom_char⇩S⇩b⇩C hom⇩C.cod_char⇩S⇩b⇩C
apply simp
by (metis (no_types, lifting) C.in_hhomI Faa'.ηε.ε.preserves_reflects_arr
hom⇩C.arr_char⇩S⇩b⇩C hom⇩C.not_arr_null hom⇩C.null_char)
show "G (F μ) ⋅⇩C C.inv (ε (C.dom μ)) = C.inv (ε (C.cod μ)) ⋅⇩C μ"
using assms 1 2 C.invert_opposite_sides_of_square by simp
qed
interpretation ε: natural_isomorphism V⇩C V⇩C GF.map C.map ε
using ε_simps ε_naturality ε_def
by unfold_locales auto
interpretation GFG: composite_functor V⇩D V⇩C V⇩C G ‹GF.map› ..
interpretation FGF: composite_functor V⇩C V⇩D V⇩D F ‹FG.map› ..
interpretation Goη: natural_transformation V⇩D V⇩C G GFG.map ‹G ∘ η›
proof -
have "G ∘ D.map = G"
using G.as_nat_trans.natural_transformation_axioms by auto
moreover have "G ∘ FG.map = GFG.map"
by auto
ultimately show "natural_transformation V⇩D V⇩C G GFG.map (G ∘ η)"
using η.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms
horizontal_composite [of V⇩D V⇩D D.map FG.map η V⇩C G G G] by simp
qed
interpretation ηoF: natural_transformation V⇩C V⇩D F FGF.map ‹η ∘ F›
using η.natural_transformation_axioms as_nat_trans.natural_transformation_axioms
horizontal_composite [of V⇩C V⇩D F F F V⇩D D.map FG.map η]
by simp
interpretation εoG: natural_transformation V⇩D V⇩C GFG.map G ‹ε ∘ G›
using ε.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms
horizontal_composite [of V⇩D V⇩C G G G V⇩C GF.map C.map ε]
by simp
interpretation Foε: natural_transformation V⇩C V⇩D FGF.map F ‹F ∘ ε›
proof -
have "F ∘ C.map = F"
using as_nat_trans.natural_transformation_axioms by auto
moreover have "F ∘ GF.map = FGF.map"
by auto
ultimately show "natural_transformation V⇩C V⇩D FGF.map F (F ∘ ε)"
using ε.natural_transformation_axioms as_nat_trans.natural_transformation_axioms
horizontal_composite [of V⇩C V⇩C GF.map C.map ε V⇩D F F F]
by simp
qed
interpretation εoG_Goη: vertical_composite V⇩D V⇩C G GFG.map G ‹G ∘ η› ‹ε ∘ G› ..
interpretation Foε_ηoF: vertical_composite V⇩C V⇩D F FGF.map F ‹η ∘ F› ‹F ∘ ε› ..
text ‹
Bijectivity results in an ordinary adjunction between the vertical categories.
›
lemma adjunction_ηε:
shows "unit_counit_adjunction V⇩C V⇩D G F η ε"
proof
show "εoG_Goη.map = G"
proof
fix ν
have "¬ D.arr ν ⟹ εoG_Goη.map ν = G ν"
unfolding εoG_Goη.map_def
by (simp add: G.extensionality)
moreover have "D.arr ν ⟹ εoG_Goη.map ν = G ν"
proof -
assume ν: "D.arr ν"
interpret hom⇩C: subcategory V⇩C ‹λμ'. «μ' : G⇩0 (src⇩D ν) →⇩C G⇩0 (trg⇩D ν)»›
using ν C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν'. «ν' : F⇩0 (G⇩0 (src⇩D ν)) →⇩D F⇩0 (G⇩0 (trg⇩D ν))»›
using D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹G⇩0 (src⇩D ν)› ‹G⇩0 (trg⇩D ν)›
using ν G⇩0_props by unfold_locales simp_all
show "εoG_Goη.map ν = G ν"
proof -
have 1: "C.arr (Faa'.G⇩1 (D.cod ν))"
using ν hom⇩D.arr_char⇩S⇩b⇩C G_def
by (metis Faa'.ηε.F.preserves_reflects_arr G_in_hom(1) hom⇩C.arr_char⇩S⇩b⇩C
hom⇩C.inclusion hom⇩D.cod_closed)
have 2: "D.arr (Faa'.η ν)"
using ν hom⇩D.arr_char⇩S⇩b⇩C
by (metis D.in_hhomI D.obj_src D.obj_trg Faa'.ηε.η.preserves_reflects_arr
G⇩0_props(1) hom⇩D.inclusion)
have 3: "src⇩C (Faa'.G⇩1 (D.cod ν)) = G⇩0 (src⇩D ν)"
using ν
by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(2))
have 4: "trg⇩C (Faa'.G⇩1 (D.cod ν)) = G⇩0 (trg⇩D ν)"
using ν
by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(3))
have 5: "hom⇩D.arr ν"
using ν hom⇩D.arr_char⇩S⇩b⇩C G⇩0_props by simp
have "εoG_Goη.map ν = ε (G (D.cod ν)) ⋅⇩C G (η ν)"
using ν εoG_Goη.map_def by simp
also have "... = Faa'.ε (Faa'.G⇩1 (D.cod ν)) ⋅⇩C Faa'.G⇩1 (Faa'.η ν)"
using ν 1 2 3 4 η_def ε_def G_def G_simps(2-3) η_simps(2-3) by auto
also have "... = hom⇩C.comp ((Faa'.ε ∘ Faa'.G⇩1) (hom⇩D.cod ν)) ((Faa'.G⇩1 ∘ Faa'.η) ν)"
proof -
have "Faa'.ε (Faa'.G⇩1 (D.cod ν)) ⋅⇩C Faa'.G⇩1 (Faa'.η ν) = Faa'.ηε.εFoFη.map ν"
proof -
have "C.seq (Faa'.ε (Faa'.G⇩1 (hom⇩D.cod ν))) (Faa'.G⇩1 (Faa'.η ν))"
proof -
have "hom⇩C.seq (Faa'.ε (Faa'.G⇩1 (hom⇩D.cod ν))) (Faa'.G⇩1 (Faa'.η ν))"
using ν 5 hom⇩C.arr_char⇩S⇩b⇩C hom⇩C.seq_char⇩S⇩b⇩C Faa'.ηε.F.preserves_arr
apply (intro hom⇩C.seqI)
apply auto
using Faa'.ηε.ε.preserves_reflects_arr hom⇩C.inclusion hom⇩D.arr_cod
by presburger
thus ?thesis
using hom⇩C.seq_char⇩S⇩b⇩C by simp
qed
moreover have "D.in_hhom (Faa'.η ν) (F⇩0 (G⇩0 (src⇩D ν))) (F⇩0 (G⇩0 (trg⇩D ν)))"
using ν 5 Faa'.ηε.η.preserves_reflects_arr hom⇩D.arrE by blast
moreover have "D.in_hhom (D.cod ν) (F⇩0 (G⇩0 (src⇩D ν))) (F⇩0 (G⇩0 (trg⇩D ν)))"
using ν 5 hom⇩D.arrE hom⇩D.cod_closed by blast
ultimately show ?thesis
using ν 5 Faa'.ηε.εFoFη.map_def hom⇩D.arr_char⇩S⇩b⇩C hom⇩C.comp_char hom⇩D.cod_char⇩S⇩b⇩C
by simp
qed
thus ?thesis
using ν 5 Faa'.ηε.εFoFη.map_def by simp
qed
also have "... = G ν"
using ν 5 G_def Faa'.ηε.triangle_F Faa'.ηε.εFoFη.map_simp_1 [of ν] by simp
finally show ?thesis by simp
qed
qed
ultimately show "εoG_Goη.map ν = G ν" by auto
qed
show "Foε_ηoF.map = F"
proof
fix μ
have "¬ C.arr μ ⟹ Foε_ηoF.map μ = F μ"
unfolding Foε_ηoF.map_def
by (simp add: extensionality)
moreover have "C.arr μ ⟹ Foε_ηoF.map μ = F μ"
proof -
assume μ: "C.arr μ"
interpret hom⇩C: subcategory V⇩C ‹λμ'. «μ' : src⇩C μ →⇩C trg⇩C μ»›
using μ C.hhom_is_subcategory by simp
interpret hom⇩D: subcategory V⇩D ‹λν. «ν : F⇩0 (src⇩C μ) →⇩D F⇩0 (trg⇩C μ)»›
using μ D.in_hhom_def D.hhom_is_subcategory by simp
interpret Faa': equivalence_pseudofunctor_at_hom
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
F Φ ‹src⇩C μ› ‹trg⇩C μ›
using μ by unfold_locales auto
show "Foε_ηoF.map μ = F μ"
proof -
have 1: "hom⇩C.arr μ"
using μ hom⇩C.arr_char⇩S⇩b⇩C by simp
have "Foε_ηoF.map μ = F (ε (C.cod μ)) ⋅⇩D η (F μ)"
using μ Foε_ηoF.map_def by simp
also have "... = Faa'.F⇩1 (Faa'.ε (C.cod μ)) ⋅⇩D Faa'.η (Faa'.F⇩1 μ)"
unfolding η_def ε_def G_def
using μ G⇩0_props Faa'.F⇩1_def
by auto
also have "... = hom⇩D.comp ((Faa'.F⇩1 ∘ Faa'.ε) (hom⇩C.cod μ)) ((Faa'.η ∘ Faa'.F⇩1) μ)"
proof -
have 2: "hom⇩C.arr (C.cod μ)"
using 1 hom⇩C.cod_char⇩S⇩b⇩C [of μ] hom⇩C.arr_cod hom⇩C.dom_char⇩S⇩b⇩C hom⇩C.cod_char⇩S⇩b⇩C
by simp
moreover have "D.seq (Faa'.F⇩1 (Faa'.ε (C.cod μ))) (Faa'.η (Faa'.F⇩1 μ))"
proof -
have "hom⇩D.seq (Faa'.F⇩1 (Faa'.ε (C.cod μ))) (Faa'.η (Faa'.F⇩1 μ))"
using μ 1 2 hom⇩D.arr_char⇩S⇩b⇩C hom⇩D.seq_char⇩S⇩b⇩C Faa'.ηε.G.preserves_arr
hom⇩C.cod_simp hom⇩C.dom_cod
apply (intro hom⇩D.seqI)
by auto metis
thus ?thesis
using hom⇩D.seq_char⇩S⇩b⇩C by simp
qed
ultimately show ?thesis
using 1 hom⇩D.comp_char hom⇩C.dom_char⇩S⇩b⇩C hom⇩C.cod_char⇩S⇩b⇩C by simp
qed
also have "... = Faa'.ηε.GεoηG.map μ"
unfolding Faa'.ηε.GεoηG.map_def
using 1 by simp
also have "... = F μ"
using μ Faa'.ηε.triangle_G hom⇩C.arr_char⇩S⇩b⇩C Faa'.ηε.εFoFη.map_def
hom⇩D.arr_char⇩S⇩b⇩C hom⇩C.comp_char hom⇩D.comp_char Faa'.F⇩1_def
by simp
finally show ?thesis by simp
qed
qed
ultimately show "Foε_ηoF.map μ = F μ" by auto
qed
qed
interpretation ηε: unit_counit_adjunction V⇩C V⇩D G F η ε
using adjunction_ηε by simp
text ‹
We now use the adjunction between the vertical categories to define the
compositors for ‹G›. Without the bijectivity assumption, we would only obtain
‹η› and ‹ε› as pseudonatural equivalences, rather than natural isomorphisms,
which would make everything more complicated.
›
definition Φ⇩G⇩0
where "Φ⇩G⇩0 f f' = C.inv (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
lemma Φ⇩G⇩0_in_hom:
assumes "D.ide f" and "D.ide f'" and "src⇩D f = trg⇩D f'"
shows "«Φ⇩G⇩0 f f' : G f ⋆⇩C G f' ⇒⇩C G (f ⋆⇩D f')»"
proof (unfold Φ⇩G⇩0_def, intro C.inv_in_hom)
show 1: "«ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')) :
G (f ⋆⇩D f') ⇒⇩C G f ⋆⇩C G f'»"
proof
show "«ε (G f ⋆⇩C G f') : G (F (G f ⋆⇩C G f')) ⇒⇩C G f ⋆⇩C G f'»"
using assms ε_in_hom by auto
show "«G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')) : G (f ⋆⇩D f') ⇒⇩C G (F (G f ⋆⇩C G f'))»"
proof -
have "«Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f') : f ⋆⇩D f' ⇒⇩D F (G f ⋆⇩C G f')»"
using assms η_in_hom
by (intro D.comp_in_homI') auto
thus ?thesis by auto
qed
qed
show "C.iso (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
proof (intro C.isos_compose)
show "C.iso (G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
proof -
have "D.iso (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f'))"
using assms η_in_hom
by (intro D.isos_compose D.seqI) auto
thus ?thesis by simp
qed
show "C.iso (ε (G f ⋆⇩C G f'))"
using assms ε_in_hom by simp
show "C.seq (ε (G f ⋆⇩C G f')) (G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
using 1 by auto
qed
qed
lemma iso_Φ⇩G⇩0:
assumes "D.ide f" and "D.ide f'" and "src⇩D f = trg⇩D f'"
shows "C.iso (Φ⇩G⇩0 f f')"
proof (unfold Φ⇩G⇩0_def, intro C.iso_inv_iso C.isos_compose)
show 1: "C.iso (G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
using assms η_in_hom D.in_hhom_def cmp_simps'(1) cmp_simps(4) by auto
show "C.iso (ε (G f ⋆⇩C G f'))"
using assms ε_in_hom by simp
show "C.seq (ε (G f ⋆⇩C G f')) (G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
using assms 1 by force
qed
lemma Φ⇩G⇩0_naturality:
assumes "D.arr ν" and "D.arr ν'" and "src⇩D ν = trg⇩D ν'"
shows "Φ⇩G⇩0 (D.cod ν) (D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν') =
G (ν ⋆⇩D ν') ⋅⇩C Φ⇩G⇩0 (D.dom ν) (D.dom ν')"
proof -
let ?X = "ε (G (D.cod ν) ⋆⇩C G (D.cod ν')) ⋅⇩C
G (Φ (G (D.cod ν), G (D.cod ν')) ⋅⇩D (η (D.cod ν) ⋆⇩D η (D.cod ν')))"
have iso_X: "C.iso ?X"
using assms C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def
by (intro C.isos_compose) auto
have "Φ⇩G⇩0 (D.cod ν) (D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν') = C.inv ?X ⋅⇩C (G ν ⋆⇩C G ν')"
unfolding Φ⇩G⇩0_def by simp
also have "... = (C.inv (G (Φ (G (D.cod ν), G (D.cod ν')) ⋅⇩D
(η (D.cod ν) ⋆⇩D η (D.cod ν')))) ⋅⇩C
C.inv (ε (G (D.cod ν) ⋆⇩C G (D.cod ν')))) ⋅⇩C
(G ν ⋆⇩C G ν')"
using assms iso_X η_in_hom ε_in_hom
by (simp add: C.inv_comp_left(1))
also have "... = (C.inv (G (Φ (G (D.cod ν), G (D.cod ν'))) ⋅⇩C
G (η (D.cod ν) ⋆⇩D η (D.cod ν'))) ⋅⇩C
C.inv (ε (G (D.cod ν) ⋆⇩C G (D.cod ν')))) ⋅⇩C
(G ν ⋆⇩C G ν')"
proof -
have "G (Φ (G (D.cod ν), G (D.cod ν')) ⋅⇩D (η (D.cod ν) ⋆⇩D η (D.cod ν'))) =
G (Φ (G (D.cod ν), G (D.cod ν'))) ⋅⇩C G (η (D.cod ν) ⋆⇩D η (D.cod ν'))"
using assms iso_X η_in_hom by fast
thus ?thesis by simp
qed
also have "... = C.inv (G (η (D.cod ν) ⋆⇩D η (D.cod ν'))) ⋅⇩C
C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) ⋅⇩C
C.inv (ε (G (D.cod ν) ⋆⇩C G (D.cod ν'))) ⋅⇩C
(G ν ⋆⇩C G ν')"
proof -
have "C.iso (G (Φ (G (D.cod ν), G (D.cod ν'))) ⋅⇩C G (η (D.cod ν) ⋆⇩D η (D.cod ν')))"
using assms iso_X η_in_hom C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def
by (intro C.isos_compose C.seqI) auto
hence "C.inv (G (Φ (G (D.cod ν), G (D.cod ν'))) ⋅⇩C
G (η (D.cod ν) ⋆⇩D η (D.cod ν')))
= C.inv (G (η (D.cod ν) ⋆⇩D η (D.cod ν'))) ⋅⇩C
C.inv (G (Φ (G (D.cod ν), G (D.cod ν'))))"
using assms η_in_hom
by (simp add: C.inv_comp_right(1))
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = G (D.inv (η (D.cod ν)) ⋆⇩D D.inv (η (D.cod ν'))) ⋅⇩C
C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) ⋅⇩C
C.inv (ε (G (D.cod ν) ⋆⇩C G (D.cod ν'))) ⋅⇩C
(G ν ⋆⇩C G ν')"
proof -
have "C.inv (G (η (D.cod ν) ⋆⇩D η (D.cod ν'))) =
G (D.inv (η (D.cod ν)) ⋆⇩D D.inv (η (D.cod ν')))"
using assms η_in_hom D.ide_cod D.iso_hcomp D.src_cod D.trg_cod G.preserves_inv
η_simps(2-3) D.inv_hcomp
by (metis D.arr_cod)
thus ?thesis by simp
qed
also have "... = G (D.inv (η (D.cod ν)) ⋆⇩D D.inv (η (D.cod ν'))) ⋅⇩C
(C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) ⋅⇩C
G (F (G ν ⋆⇩C G ν'))) ⋅⇩C
C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')))"
proof -
have "C.inv (ε (G (D.cod ν) ⋆⇩C G (D.cod ν'))) ⋅⇩C (G ν ⋆⇩C G ν')
= G (F (G ν ⋆⇩C G ν')) ⋅⇩C C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')))"
proof -
have "G (D.dom ν) ⋆⇩C G (D.dom ν') = C.dom (G ν ⋆⇩C G ν')"
by (simp add: assms)
thus ?thesis
using assms C.hcomp_simps(4) C.hseqI' G.preserves_cod G.preserves_hseq
ε_naturality(3)
by presburger
qed
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = (G (D.inv (η (D.cod ν)) ⋆⇩D D.inv (η (D.cod ν'))) ⋅⇩C
G (F (G ν) ⋆⇩D F (G ν'))) ⋅⇩C
G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) ⋅⇩C
C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')))"
proof -
have "C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) ⋅⇩C G (F (G ν ⋆⇩C G ν')) =
G (D.inv (Φ (C.cod (G ν), C.cod (G ν'))) ⋅⇩D F (G ν ⋆⇩C G ν'))"
using assms by (simp add: G.preserves_inv)
also have "... = G ((F (G ν) ⋆⇩D F (G ν')) ⋅⇩D D.inv (Φ (C.dom (G ν), C.dom (G ν'))))"
using assms C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def
Φ.inv_naturality [of "(G ν, G ν')"]
by simp
also have "... = G (F (G ν) ⋆⇩D F (G ν')) ⋅⇩C
G (D.inv (Φ (C.dom (G ν), C.dom (G ν'))))"
using assms by simp
finally show ?thesis
using C.comp_assoc by simp
qed
also have "... = G (ν ⋆⇩D ν') ⋅⇩C
G (D.inv (η (D.dom ν)) ⋆⇩D D.inv (η (D.dom ν'))) ⋅⇩C
G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) ⋅⇩C
C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')))"
proof -
have "G (D.inv (η (D.cod ν)) ⋆⇩D D.inv (η (D.cod ν'))) ⋅⇩C G (F (G ν) ⋆⇩D F (G ν')) =
G ((D.inv (η (D.cod ν)) ⋆⇩D D.inv (η (D.cod ν'))) ⋅⇩D (F (G ν) ⋆⇩D F (G ν')))"
using assms by simp
also have "... = G (D.inv (η (D.cod ν)) ⋅⇩D F (G ν)
⋆⇩D D.inv (η (D.cod ν')) ⋅⇩D F (G ν'))"
using assms D.interchange by simp
also have "... = G (ν ⋅⇩D D.inv (η (D.dom ν)) ⋆⇩D ν' ⋅⇩D D.inv (η (D.dom ν')))"
using assms η_naturality by simp
also have "... = G ((ν ⋆⇩D ν') ⋅⇩D (D.inv (η (D.dom ν)) ⋆⇩D D.inv (η (D.dom ν'))))"
using assms D.interchange by simp
also have "... = G (ν ⋆⇩D ν') ⋅⇩C G (D.inv (η (D.dom ν)) ⋆⇩D D.inv (η (D.dom ν')))"
using assms by simp
finally show ?thesis
using C.comp_assoc by simp
qed
also have "... = G (ν ⋆⇩D ν') ⋅⇩C Φ⇩G⇩0 (D.dom ν) (D.dom ν')"
proof -
have "G (D.inv (η (D.dom ν)) ⋆⇩D D.inv (η (D.dom ν'))) ⋅⇩C
G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) ⋅⇩C
C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')))
= C.inv (G (η (D.dom ν) ⋆⇩D η (D.dom ν'))) ⋅⇩C
C.inv (G (Φ (C.dom (G ν), C.dom (G ν')))) ⋅⇩C
C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')))"
proof -
have "G (D.inv (η (D.dom ν)) ⋆⇩D D.inv (η (D.dom ν'))) =
C.inv (G (η (D.dom ν) ⋆⇩D η (D.dom ν')))"
using assms cmp_components_are_iso
by (metis D.arr_dom D.ide_dom D.inv_hcomp D.iso_hcomp D.src_dom D.trg_dom
G.preserves_inv η_simps(2) η_simps(3) η_simps(6))
moreover have "G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) =
C.inv (G (Φ (C.dom (G ν), C.dom (G ν'))))"
using assms cmp_components_are_iso
by (simp add: G.preserves_inv)
ultimately show ?thesis by simp
qed
also have "... = C.inv (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')) ⋅⇩C
G (Φ (C.dom (G ν), C.dom (G ν'))) ⋅⇩C
G (η (D.dom ν) ⋆⇩D η (D.dom ν')))"
proof -
have "C.iso (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')) ⋅⇩C
G (Φ (C.dom (G ν), C.dom (G ν'))) ⋅⇩C
G (η (D.dom ν) ⋆⇩D η (D.dom ν')))"
using assms η_in_hom ε_in_hom cmp_simps'(1,4) C.VV.cod_char⇩S⇩b⇩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 (ε (G (D.dom ν) ⋆⇩C G (D.dom ν')) ⋅⇩C
G (Φ (C.dom (G ν), C.dom (G ν')) ⋅⇩D
(η (D.dom ν) ⋆⇩D η (D.dom ν'))))"
proof -
have "G (Φ (C.dom (G ν), C.dom (G ν'))) ⋅⇩C G (η (D.dom ν) ⋆⇩D η (D.dom ν')) =
G (Φ (C.dom (G ν), C.dom (G ν')) ⋅⇩D (η (D.dom ν) ⋆⇩D η (D.dom ν')))"
using assms cmp_simps'(1,4) by auto
thus ?thesis by simp
qed
also have "... = Φ⇩G⇩0 (D.dom ν) (D.dom ν')"
unfolding Φ⇩G⇩0_def
using assms by simp
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
interpretation Φ⇩G: transformation_by_components D.VV.comp V⇩C H⇩DoGG.map GoH⇩D.map
‹λfg. Φ⇩G⇩0 (fst fg) (snd fg)›
using D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C Φ⇩G⇩0_in_hom Φ⇩G⇩0_naturality G.FF_def
D.VV.dom_char⇩S⇩b⇩C D.VV.cod_char⇩S⇩b⇩C
by unfold_locales auto
interpretation Φ⇩G: natural_isomorphism D.VV.comp V⇩C H⇩DoGG.map GoH⇩D.map Φ⇩G.map
using Φ⇩G.map_simp_ide D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C iso_Φ⇩G⇩0 by unfold_locales simp
abbreviation Φ⇩G
where "Φ⇩G ≡ Φ⇩G.map"
lemma Φ⇩G_in_hom [intro]:
assumes "D.arr ν" and "D.arr ν'" and "src⇩D ν = trg⇩D ν'"
shows "C.in_hhom (Φ⇩G.map (ν, ν')) (src⇩C (G (D.dom ν'))) (trg⇩C (G (D.cod ν)))"
and "«Φ⇩G.map (ν, ν') : G (D.dom ν) ⋆⇩C G (D.dom ν') ⇒⇩C G (D.cod ν ⋆⇩D D.cod ν')»"
proof -
show "«Φ⇩G.map (ν, ν') : G (D.dom ν) ⋆⇩C G (D.dom ν') ⇒⇩C G (D.cod ν ⋆⇩D D.cod ν')»"
proof -
have "Φ⇩G.map (ν, ν') = Φ⇩G⇩0 (D.cod ν) (D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν')"
using assms D.VV.arr_char⇩S⇩b⇩C Φ⇩G.map_def Φ⇩G⇩0_in_hom G.FF_def D.VV.cod_char⇩S⇩b⇩C
by simp
moreover have "«Φ⇩G⇩0 (D.cod ν) (D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν') :
G (D.dom ν) ⋆⇩C G (D.dom ν') ⇒⇩C G (D.cod ν ⋆⇩D D.cod ν')»"
using assms Φ⇩G⇩0_in_hom by (intro C.comp_in_homI) auto
ultimately show ?thesis by simp
qed
thus "C.in_hhom (Φ⇩G.map (ν, ν')) (src⇩C (G (D.dom ν'))) (trg⇩C (G (D.cod ν)))"
using assms C.vconn_implies_hpar(1-2) by auto
qed
lemma Φ⇩G_simps [simp]:
assumes "D.arr ν" and "D.arr ν'" and "src⇩D ν = trg⇩D ν'"
shows "C.arr (Φ⇩G.map (ν, ν'))"
and "src⇩C (Φ⇩G.map (ν, ν')) = src⇩C (G (D.dom ν'))"
and "trg⇩C (Φ⇩G.map (ν, ν')) = trg⇩C (G (D.cod ν))"
and "C.dom (Φ⇩G.map (ν, ν')) = G (D.dom ν) ⋆⇩C G (D.dom ν')"
and "C.cod (Φ⇩G.map (ν, ν')) = G (D.cod ν ⋆⇩D D.cod ν')"
using assms Φ⇩G_in_hom
apply auto
by fast+
lemma Φ⇩G_map_simp_ide:
assumes "D.ide f" and "D.ide f'" and "src⇩D f = trg⇩D f'"
shows "Φ⇩G.map (f, f') = G (D.inv (η f) ⋆⇩D D.inv (η f')) ⋅⇩C G (D.inv (Φ (G f, G f'))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G f'))"
proof -
have "Φ⇩G.map (f, f') = C.inv (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))"
using assms Φ⇩G.map_simp_ide D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C Φ⇩G⇩0_def G.FF_def by auto
also have "... = C.inv (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f')) ⋅⇩C G (η f ⋆⇩D η f'))"
using assms D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C cmp_simps(1,4) by simp
also have "... = C.inv (G (η f ⋆⇩D η f')) ⋅⇩C C.inv (G (Φ (G f, G f'))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G f'))"
proof -
have "C.iso (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f')) ⋅⇩C G (η f ⋆⇩D η f'))"
using assms C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩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 (η f) ⋆⇩D D.inv (η f')) ⋅⇩C G (D.inv (Φ (G f, G f'))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G f'))"
using assms cmp_components_are_iso D.ideD(1) D.inv_hcomp D.iso_hcomp
G.preserves_ide G.preserves_inv G_simps(2-3) η_simps(2-3,6)
by metis
finally show ?thesis by blast
qed
lemma η_hcomp:
assumes "D.ide f" and "D.ide f'" and "src⇩D f = trg⇩D f'"
shows "η (f ⋆⇩D f') = F (Φ⇩G.map (f, f')) ⋅⇩D Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')"
proof -
have 1: "«F (Φ⇩G.map (f, f')) ⋅⇩D Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f') :
f ⋆⇩D f' ⇒⇩D F (G (f ⋆⇩D f'))»"
using assms by fastforce
have 2: "«ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f')) ⋅⇩C G (η f ⋆⇩D η f') :
G (f ⋆⇩D f') ⇒⇩C G f ⋆⇩C G f'»"
using assms G.preserves_hom cmp_in_hom(2)
by (intro C.comp_in_homI) auto
have "F (Φ⇩G.map (f, f')) ⋅⇩D Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')
= F (C.inv (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')))) ⋅⇩D
Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')"
using assms Φ⇩G.map_simp_ide D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C Φ⇩G⇩0_def by simp
also have "... = F (C.inv (G (η f ⋆⇩D η f')) ⋅⇩C C.inv (G (Φ (G f, G f'))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G f'))) ⋅⇩D
Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')"
proof -
have "C.inv (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f'))) =
C.inv (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f')) ⋅⇩C G (η f ⋆⇩D η f'))"
using assms 1 by force
also have "... = C.inv (G (η f ⋆⇩D η f')) ⋅⇩C C.inv (G (Φ (G f, G f'))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G f'))"
proof -
have "C.iso (ε (G f ⋆⇩C G f') ⋅⇩C G (Φ (G f, G f')) ⋅⇩C G (η f ⋆⇩D η f'))"
using assms 2 by (intro C.isos_compose) auto
thus ?thesis
using assms 1 C.inv_comp_left C.comp_assoc by simp
qed
finally show ?thesis by simp
qed
also have "... = F (C.inv (G (η f ⋆⇩D η f'))) ⋅⇩D
F (C.inv (G (Φ (G f, G f')))) ⋅⇩D F (C.inv (ε (G f ⋆⇩C G f'))) ⋅⇩D
Φ (G f, G f') ⋅⇩D (η f ⋆⇩D η f')"
proof -
have "C.arr ((C.inv (G (η f ⋆⇩D η f')) ⋅⇩C C.inv (G (Φ (G f, G f'))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G f'))))"
using assms 1 2 cmp_components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩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 (η f ⋆⇩D η f'))) ⋅⇩D
D.inv (F (G (Φ (G f, G f')))) ⋅⇩D
D.inv (F (ε (G f ⋆⇩C G f'))) ⋅⇩D
Φ (G f, G f') ⋅⇩D
(η f ⋆⇩D η f')"
using assms by (simp add: preserves_inv)
also have "... = D.inv ((η (F (G f) ⋆⇩D F (G f'))) ⋅⇩D (η f ⋆⇩D η f') ⋅⇩D
D.inv (η (f ⋆⇩D f'))) ⋅⇩D
D.inv (η (F (G f ⋆⇩C G f')) ⋅⇩D Φ (G f, G f') ⋅⇩D
D.inv (η (F (G f) ⋆⇩D F (G f')))) ⋅⇩D
D.inv (F (ε (G f ⋆⇩C G f'))) ⋅⇩D
Φ (G f, G f') ⋅⇩D
(η f ⋆⇩D η f')"
proof -
have *: "⋀ν. D.arr ν ⟹ η (D.cod ν) ⋅⇩D ν ⋅⇩D D.inv (η (D.dom ν)) = F (G ν)"
by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
η_naturality(1-2) η_simps(1,6))
have "F (G (η f ⋆⇩D η f')) =
η (F (G f) ⋆⇩D F (G f')) ⋅⇩D (η f ⋆⇩D η f') ⋅⇩D D.inv (η (f ⋆⇩D f'))"
using assms * [of "η f ⋆⇩D η f'"] by simp
moreover have "F (G (Φ (G f, G f'))) =
η (F (G f ⋆⇩C G f')) ⋅⇩D Φ (G f, G f') ⋅⇩D D.inv (η (F (G f) ⋆⇩D F (G f')))"
using assms 1 * [of "Φ (G f, G f')"] cmp_simps(5) by fastforce
ultimately show ?thesis by simp
qed
also have "... = η (f ⋆⇩D f') ⋅⇩D
D.inv (η f ⋆⇩D η f') ⋅⇩D
((D.inv (η (F (G f) ⋆⇩D F (G f')))) ⋅⇩D (η (F (G f) ⋆⇩D F (G f')))) ⋅⇩D
(D.inv (Φ (G f, G f')) ⋅⇩D
((D.inv (η (F (G f ⋆⇩C G f')))) ⋅⇩D D.inv (F (ε (G f ⋆⇩C G f')))) ⋅⇩D
Φ (G f, G f')) ⋅⇩D
(η f ⋆⇩D η f')"
proof -
have "D.inv ((η (F (G f) ⋆⇩D F (G f'))) ⋅⇩D (η f ⋆⇩D η f') ⋅⇩D D.inv (η (f ⋆⇩D f')))
= η (f ⋆⇩D f') ⋅⇩D D.inv (η f ⋆⇩D η f') ⋅⇩D D.inv (η (F (G f) ⋆⇩D F (G f')))"
proof -
have "D.iso (η (F (G f) ⋆⇩D F (G f')) ⋅⇩D (η f ⋆⇩D η f') ⋅⇩D D.inv (η (f ⋆⇩D f')))"
using assms by (intro D.isos_compose) auto
thus ?thesis
using assms D.inv_comp_left D.comp_assoc by simp
qed
moreover have
"D.inv (η (F (G f ⋆⇩C G f')) ⋅⇩D Φ (G f, G f') ⋅⇩D D.inv (η (F (G f) ⋆⇩D F (G f')))) =
η (F (G f) ⋆⇩D F (G f')) ⋅⇩D D.inv (Φ (G f, G f')) ⋅⇩D D.inv (η (F (G f ⋆⇩C G f')))"
proof -
have "D.iso (η (F (G f ⋆⇩C G f')) ⋅⇩D Φ (G f, G f') ⋅⇩D D.inv (η (F (G f) ⋆⇩D F (G f'))))"
proof -
have 1: "D.seq (Φ (G f, G f')) (D.inv (η (F (G f) ⋆⇩D F (G f'))))"
using assms by (intro D.seqI) auto
moreover have "D.seq (η (F (G f ⋆⇩C G f')))
(Φ (G f, G f') ⋅⇩D D.inv (η (F (G f) ⋆⇩D F (G f'))))"
using assms 1 by auto
ultimately show ?thesis
using assms cmp_components_are_iso η_in_hom
by (intro D.isos_compose) auto
qed
thus ?thesis
using assms D.comp_inv_arr' D.comp_assoc D.inv_comp_left cmp_components_are_iso
by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D f') ⋅⇩D
D.inv (η f ⋆⇩D η f') ⋅⇩D
(D.inv (η (F (G f) ⋆⇩D F (G f'))) ⋅⇩D (η (F (G f) ⋆⇩D F (G f')))) ⋅⇩D
(D.inv (Φ (G f, G f')) ⋅⇩D Φ (G f, G f')) ⋅⇩D
(η f ⋆⇩D η f')"
proof -
have "(D.inv (η (F (G f ⋆⇩C G f')))) ⋅⇩D D.inv (F (ε (G f ⋆⇩C G f'))) ⋅⇩D Φ (G f, G f')
= F (G f ⋆⇩C G f') ⋅⇩D Φ (G f, G f')"
proof -
have "(D.inv (η (F (G f ⋆⇩C G f')))) ⋅⇩D D.inv (F (ε (G f ⋆⇩C G f')))
= D.inv (F (ε (G f ⋆⇩C G f')) ⋅⇩D η (F (G f ⋆⇩C G f')))"
using assms by (simp add: D.inv_comp)
also have "... = F (G f ⋆⇩C G f')"
using assms ηε.triangle_G Foε_ηoF.map_simp_ide by fastforce
finally show ?thesis using D.comp_assoc by metis
qed
also have "... = Φ (G f, G f')"
using assms D.comp_cod_arr cmp_in_hom(2) [of "G f" "G f'"] by auto
finally have "(D.inv (η (F (G f ⋆⇩C G f')))) ⋅⇩D D.inv (F (ε (G f ⋆⇩C G f'))) ⋅⇩D
Φ (G f, G f')
= Φ (G f, G f')"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D f') ⋅⇩D
D.inv (η f ⋆⇩D η f') ⋅⇩D
(D.inv (η (F (G f) ⋆⇩D F (G f'))) ⋅⇩D (η (F (G f) ⋆⇩D F (G f')))) ⋅⇩D
(η f ⋆⇩D η f')"
using assms D.comp_cod_arr D.comp_inv_arr' cmp_components_are_iso by simp
also have "... = η (f ⋆⇩D f') ⋅⇩D D.inv (η f ⋆⇩D η f') ⋅⇩D (η f ⋆⇩D η f')"
using assms η_in_hom D.comp_inv_arr' D.comp_cod_arr by simp
also have "... = η (f ⋆⇩D f')"
using assms D.comp_inv_arr' [of "η f ⋆⇩D η f'"] D.comp_arr_dom by simp
finally show ?thesis by simp
qed
lemma ε_hcomp:
assumes "C.ide g" and "C.ide g'" and "src⇩C g = trg⇩C g'"
shows "ε (g ⋆⇩C g') = (ε g ⋆⇩C ε g') ⋅⇩C C.inv (Φ⇩G.map (F g, F g')) ⋅⇩C C.inv (G (Φ (g, g')))"
proof -
have 1: "«ε (G (F g) ⋆⇩C G (F g')) ⋅⇩C G (Φ (G (F g), G (F g')) ⋅⇩D (η (F g) ⋆⇩D η (F g'))) :
G (F g ⋆⇩D F g') ⇒⇩C G (F g) ⋆⇩C G (F g')»"
using assms by fastforce
have "(ε g ⋆⇩C ε g') ⋅⇩C C.inv (Φ⇩G.map (F g, F g')) ⋅⇩C C.inv (G (Φ (g, g')))
= (ε g ⋆⇩C ε g') ⋅⇩C
C.inv (C.inv (ε (G (F g) ⋆⇩C G (F g')) ⋅⇩C G (Φ (G (F g), G (F g')) ⋅⇩D
(η (F g) ⋆⇩D η (F g'))))) ⋅⇩C
C.inv (G (Φ (g, g')))"
using assms Φ⇩G.map_simp_ide D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C Φ⇩G⇩0_def by simp
also have "... = ((ε g ⋆⇩C ε g') ⋅⇩C ε (G (F g) ⋆⇩C G (F g'))) ⋅⇩C
G (Φ (G (F g), G (F g')) ⋅⇩D (η (F g) ⋆⇩D η (F g'))) ⋅⇩C
C.inv (G (Φ (g, g')))"
proof -
have "C.iso (ε (G (F g) ⋆⇩C G (F g')) ⋅⇩C G (Φ (G (F g), G (F g')) ⋅⇩D
(η (F g) ⋆⇩D η (F g'))))"
using assms 1
by (intro C.isos_compose D.isos_compose G.preserves_iso) auto
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = ε (g ⋆⇩C g') ⋅⇩C
(G (F (ε g ⋆⇩C ε g')) ⋅⇩C G (Φ (G (F g), G (F g')) ⋅⇩D
(η (F g) ⋆⇩D η (F g')))) ⋅⇩C
C.inv (G (Φ (g, g')))"
proof -
have "(ε g ⋆⇩C ε g') ⋅⇩C ε (G (F g) ⋆⇩C G (F g')) = ε (g ⋆⇩C g') ⋅⇩C G (F (ε g ⋆⇩C ε g'))"
using assms ε_naturality [of "ε g ⋆⇩C ε g'"] by simp
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = ε (g ⋆⇩C g') ⋅⇩C
G ((F (ε g ⋆⇩C ε g') ⋅⇩D Φ (G (F g), G (F g'))) ⋅⇩D (η (F g) ⋆⇩D η (F g'))) ⋅⇩C
C.inv (G (Φ (g, g')))"
using assms C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C FF_def D.comp_assoc by auto
also have "... = ε (g ⋆⇩C g') ⋅⇩C
G (Φ (g, g') ⋅⇩D (F (ε g) ⋆⇩D F (ε g')) ⋅⇩D (η (F g) ⋆⇩D η (F g'))) ⋅⇩C
C.inv (G (Φ (g, g')))"
proof -
have "F (ε g ⋆⇩C ε g') ⋅⇩D Φ (G (F g), G (F g')) = Φ (g, g') ⋅⇩D (F (ε g) ⋆⇩D F (ε g'))"
using assms C.VV.arr_char⇩S⇩b⇩C D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C FF_def
Φ.naturality [of "(ε g, ε g')"] C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ε (g ⋆⇩C g') ⋅⇩C
G (Φ (g, g') ⋅⇩D (F (ε g) ⋅⇩D η (F g) ⋆⇩D F (ε g') ⋅⇩D η (F g'))) ⋅⇩C
C.inv (G (Φ (g, g')))"
using assms D.interchange by simp
also have "... = ε (g ⋆⇩C g') ⋅⇩C G (Φ (g, g') ⋅⇩D (F g ⋆⇩D F g')) ⋅⇩C C.inv (G (Φ (g, g')))"
using assms ηε.triangle_G Foε_ηoF.map_def [of g] Foε_ηoF.map_def [of g'] by simp
also have "... = ε (g ⋆⇩C g') ⋅⇩C G (Φ (g, g')) ⋅⇩C C.inv (G (Φ (g, g')))"
using assms D.comp_arr_dom cmp_simps(1) cmp_simps(4) by auto
also have "... = ε (g ⋆⇩C g')"
using assms D.comp_arr_dom C.comp_arr_inv' C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C
cmp_components_are_iso C.comp_arr_dom
by simp
finally show ?thesis by simp
qed
lemma G_preserves_hcomp:
assumes "D.hseq ν ν'"
shows "G (ν ⋆⇩D ν') = Φ⇩G.map (D.cod ν, D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν') ⋅⇩C
C.inv (Φ⇩G.map (D.dom ν, D.dom ν'))"
proof -
have "G (ν ⋆⇩D ν') ⋅⇩C Φ⇩G.map (D.dom ν, D.dom ν') =
Φ⇩G.map (D.cod ν, D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν')"
using assms Φ⇩G.naturality D.VV.arr_char⇩S⇩b⇩C D.VV.cod_char⇩S⇩b⇩C G.FF_def D.VV.dom_char⇩S⇩b⇩C by auto
moreover have "C.seq (Φ⇩G.map (D.cod ν, D.cod ν')) (G ν ⋆⇩C G ν')"
proof
show "«G ν ⋆⇩C G ν' :
G (D.dom ν) ⋆⇩C G (D.dom ν') ⇒⇩C G (D.cod ν) ⋆⇩C G (D.cod ν')»"
using assms G_in_hom(2) C.hcomp_in_vhom by auto
show "«Φ⇩G.map (D.cod ν, D.cod ν') :
G (D.cod ν) ⋆⇩C G (D.cod ν') ⇒⇩C G (D.cod ν ⋆⇩D D.cod ν')»"
using assms Φ⇩G⇩0_in_hom Φ⇩G.map_simp_ide D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C by auto
qed
moreover have "C.iso (Φ⇩G.map (D.dom ν, D.dom ν'))"
using assms Φ⇩G.components_are_iso D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C by auto
ultimately show ?thesis
using assms Φ⇩G.components_are_iso C.comp_assoc
C.invert_side_of_triangle(2)
[of "Φ⇩G.map (D.cod ν, D.cod ν') ⋅⇩C (G ν ⋆⇩C G ν')"
"G (ν ⋆⇩D ν')" "Φ⇩G.map (D.dom ν, D.dom ν')"]
by simp
qed
lemma coherence_LHS:
assumes "D.ide f" and "D.ide g" and "D.ide h"
and "src⇩D f = trg⇩D g" and "src⇩D g = trg⇩D h"
shows "F (G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C (Φ⇩G.map (f, g) ⋆⇩C G h))
= (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D (D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
𝖺⇩D[F (G f), F (G g), F (G h)] ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "F (G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C (Φ⇩G.map (f, g) ⋆⇩C G h))
= F (G 𝖺⇩D[f, g, h] ⋅⇩C
(G (D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩C
G (D.inv (Φ (G (f ⋆⇩D g), G h))) ⋅⇩C
C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h)) ⋅⇩C
(G (D.inv (η f) ⋆⇩D D.inv (η g)) ⋅⇩C
G (D.inv (Φ (G f, G g))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G g)) ⋆⇩C G h))"
using assms Φ⇩G_map_simp_ide C.comp_assoc by simp
also have "... = F (G 𝖺⇩D[f, g, h] ⋅⇩C
(G (D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩C
G (D.inv (Φ (G (f ⋆⇩D g), G h))) ⋅⇩C
C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h)) ⋅⇩C
(G (D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩C G h) ⋅⇩C
(G (D.inv (Φ (G f, G g))) ⋆⇩C G h) ⋅⇩C
(C.inv (ε (G f ⋆⇩C G g)) ⋆⇩C G h))"
using assms C.whisker_right D.comp_assoc by simp
also have "... = F (G 𝖺⇩D[f, g, h]) ⋅⇩D
F (G (D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
F (G (D.inv (Φ (G (f ⋆⇩D g), G h)))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
F (G (D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩C G h) ⋅⇩D
F (G (D.inv (Φ (G f, G g))) ⋆⇩C G h) ⋅⇩D
F (C.inv (ε (G f ⋆⇩C G g)) ⋆⇩C G h)"
proof -
have "C.arr (G 𝖺⇩D[f, g, h] ⋅⇩C
(G (D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩C
G (D.inv (Φ (G (f ⋆⇩D g), G h))) ⋅⇩C
C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h)) ⋅⇩C
(G (D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩C G h) ⋅⇩C
(G (D.inv (Φ (G f, G g))) ⋆⇩C G h) ⋅⇩C
(C.inv (ε (G f ⋆⇩C G g)) ⋆⇩C G h))"
using assms Φ⇩G_simps G.FF_def G⇩0_props by auto
thus ?thesis
by (metis C.seqE as_nat_trans.preserves_comp_2)
qed
also have "... = F (G 𝖺⇩D[f, g, h]) ⋅⇩D
F (G (D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
F (G (D.inv (Φ (G (f ⋆⇩D g), G h)))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
(Φ (G (f ⋆⇩D g), G h) ⋅⇩D
(F (G (D.inv (η f) ⋆⇩D D.inv (η g))) ⋆⇩D F (G h)) ⋅⇩D
((D.inv (Φ (G (F (G f) ⋆⇩D F (G g)), G h))) ⋅⇩D
(Φ (G (F (G f) ⋆⇩D F (G g)), G h)) ⋅⇩D
(F (G (D.inv (Φ (G f, G g)))) ⋆⇩D F (G h))) ⋅⇩D
((D.inv (Φ (G (F (G f ⋆⇩C G g)), G h))) ⋅⇩D
(Φ (G (F (G f ⋆⇩C G g)), G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h))) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h)))"
using assms G⇩0_props preserves_hcomp FF_def D.comp_assoc by auto
also have "... = F (G 𝖺⇩D[f, g, h]) ⋅⇩D
F (G (D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
F (G (D.inv (Φ (G (f ⋆⇩D g), G h)))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
(Φ (G (f ⋆⇩D g), G h) ⋅⇩D
(F (G (D.inv (η f) ⋆⇩D D.inv (η g))) ⋆⇩D F (G h)) ⋅⇩D
(F (G (D.inv (Φ (G f, G g)))) ⋆⇩D F (G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h)))"
proof -
have "((D.inv (Φ (G (F (G f) ⋆⇩D F (G g)), G h))) ⋅⇩D
(Φ (G (F (G f) ⋆⇩D F (G g)), G h))) ⋅⇩D
(F (G (D.inv (Φ (G f, G g)))) ⋆⇩D F (G h))
= F (G (D.inv (Φ (G f, G g)))) ⋆⇩D F (G h)"
using assms Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C G⇩0_props FF_def
D.comp_inv_arr' D.comp_cod_arr
by simp
moreover have "(D.inv (Φ (G (F (G f ⋆⇩C G g)), G h)) ⋅⇩D
Φ (G (F (G f ⋆⇩C G g)), G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h)) =
F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h)"
using assms D.comp_cod_arr Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C
G⇩0_props FF_def D.comp_inv_arr'
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
𝖺⇩D[f, g, h] ⋅⇩D
((D.inv (η ((f ⋆⇩D g) ⋆⇩D h))) ⋅⇩D
(η ((f ⋆⇩D g) ⋆⇩D h)) ⋅⇩D
(D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
(((D.inv (η (F (G (f ⋆⇩D g)) ⋆⇩D F (G h)))) ⋅⇩D
(η (F (G (f ⋆⇩D g)) ⋆⇩D F (G h)))) ⋅⇩D
D.inv (Φ (G (f ⋆⇩D g), G h))) ⋅⇩D
D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
(Φ (G (f ⋆⇩D g), G h) ⋅⇩D
((η (f ⋆⇩D g) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g)) ⋅⇩D
D.inv (η (F (G f) ⋆⇩D F (G g)))) ⋆⇩D F (G h)) ⋅⇩D
((η (F (G f) ⋆⇩D F (G g)) ⋅⇩D
D.inv (Φ (G f, G g)) ⋅⇩D
D.inv (η (F (G f ⋆⇩C G g)))) ⋆⇩D F (G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h)))"
proof -
have "⋀ν. D.arr ν ⟹ F (G ν) = η (D.cod ν) ⋅⇩D ν ⋅⇩D D.inv (η (D.dom ν))"
by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
η_naturality(1-2) η_simps(1,6))
thus ?thesis
using assms D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
𝖺⇩D[f, g, h] ⋅⇩D
(D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h)) ⋅⇩D
D.inv (Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
(Φ (G (f ⋆⇩D g), G h) ⋅⇩D
((η (f ⋆⇩D g) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g)) ⋅⇩D
D.inv (η (F (G f) ⋆⇩D F (G g))) ⋆⇩D F (G h)) ⋅⇩D
(η (F (G f) ⋆⇩D F (G g)) ⋅⇩D
D.inv (Φ (G f, G g)) ⋅⇩D
D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h))) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h)))"
proof -
have "(D.inv (η ((f ⋆⇩D g) ⋆⇩D h)) ⋅⇩D η ((f ⋆⇩D g) ⋆⇩D h)) ⋅⇩D
(D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))
= D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h)"
using assms D.comp_inv_arr' D.comp_cod_arr η_in_hom by simp
moreover have "((D.inv (η (F (G (f ⋆⇩D g)) ⋆⇩D F (G h)))) ⋅⇩D
η (F (G (f ⋆⇩D g)) ⋆⇩D F (G h))) ⋅⇩D
D.inv (Φ (G (f ⋆⇩D g), G h))
= D.inv (Φ (G (f ⋆⇩D g), G h))"
using assms D.comp_inv_arr' D.comp_cod_arr by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
𝖺⇩D[f, g, h] ⋅⇩D
((D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
(D.inv (Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h)))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D
(Φ (G (f ⋆⇩D g), G h) ⋅⇩D
(η (f ⋆⇩D g) ⋆⇩D F (G h)) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h)) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
(((D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h))) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))))"
proof -
have "(η (f ⋆⇩D g) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g)) ⋅⇩D
D.inv (η (F (G f) ⋆⇩D F (G g))) ⋆⇩D F (G h)) ⋅⇩D
(η (F (G f) ⋆⇩D F (G g)) ⋅⇩D
D.inv (Φ (G f, G g)) ⋅⇩D
D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h))
= (η (f ⋆⇩D g) ⋆⇩D F (G h)) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h)) ⋅⇩D
(((D.inv (η (F (G f) ⋆⇩D F (G g))) ⋆⇩D F (G h)) ⋅⇩D
(η (F (G f) ⋆⇩D F (G g)) ⋆⇩D F (G h))) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h))) ⋅⇩D
(D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h))"
using assms D.comp_assoc D.whisker_right by simp
also have "... = (η (f ⋆⇩D g) ⋆⇩D F (G h)) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h)) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
(D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h))"
proof -
have "((D.inv (η (F (G f) ⋆⇩D F (G g))) ⋆⇩D F (G h)) ⋅⇩D
(η (F (G f) ⋆⇩D F (G g)) ⋆⇩D F (G h))) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h))
= D.inv (Φ (G f, G g)) ⋆⇩D F (G h)"
using assms D.comp_inv_arr' D.comp_cod_arr
D.interchange [of "D.inv (η (F (G f) ⋆⇩D F (G g)))" "η (F (G f) ⋆⇩D F (G g))"
"F (G h)" "F (G h)"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
𝖺⇩D[f, g, h] ⋅⇩D
((D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
(D.inv (Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
((D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h)))) ⋅⇩D
F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h)))) ⋅⇩D
Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
(η (f ⋆⇩D g) ⋆⇩D F (G h)) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h)) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "((D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h))) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))
= D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "((D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
(F (C.inv (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h))) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))
= ((D.inv (η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
(D.inv (F (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h))) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
using assms by (simp add: preserves_inv)
also have "... = (D.inv (η (F (G f ⋆⇩C G g))) ⋅⇩D
D.inv (F (ε (G f ⋆⇩C G g))) ⋆⇩D F (G h) ⋅⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
using assms Φ.components_are_iso
D.interchange [of "D.inv (η (F (G f ⋆⇩C G g)))" "D.inv (F (ε (G f ⋆⇩C G g)))"
"F (G h)" "F (G h)"]
by simp
also have "... = (D.inv (F (ε (G f ⋆⇩C G g)) ⋅⇩D η (F (G f ⋆⇩C G g))) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "D.iso (F (ε (G f ⋆⇩C G g)) ⋅⇩D η (F (G f ⋆⇩C G g)))"
using assms by auto
hence "D.inv (η (F (G f ⋆⇩C G g))) ⋅⇩D D.inv (F (ε (G f ⋆⇩C G g))) =
D.inv (F (ε (G f ⋆⇩C G g)) ⋅⇩D η (F (G f ⋆⇩C G g)))"
using assms by (simp add: D.inv_comp_right(1))
thus ?thesis
using assms by simp
qed
also have "... = (F (G f ⋆⇩C G g) ⋆⇩D F (G h)) ⋅⇩D D.inv (Φ (G f ⋆⇩C G g, G h))"
using assms ηε.triangle_G Foε_ηoF.map_def [of "G f ⋆⇩C G g"] by simp
also have "... = D.inv (Φ (G f ⋆⇩C G g, G h))"
using assms D.comp_cod_arr by simp
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
𝖺⇩D[f, g, h] ⋅⇩D
((D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h))) ⋅⇩D
((D.inv (Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
(η (f ⋆⇩D g) ⋆⇩D F (G h))) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h)) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "((D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))))) ⋅⇩D
(Φ (G (f ⋆⇩D g), G h))
= Φ (G (f ⋆⇩D g), G h)"
proof -
have "D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D F (C.inv (ε (G (f ⋆⇩D g) ⋆⇩C G h))) =
D.inv (η (F (G (f ⋆⇩D g) ⋆⇩C G h))) ⋅⇩D D.inv (F (ε (G (f ⋆⇩D g) ⋆⇩C G h)))"
using assms preserves_inv by simp
also have "... = D.inv (F (ε (G (f ⋆⇩D g) ⋆⇩C G h)) ⋅⇩D η (F (G (f ⋆⇩D g) ⋆⇩C G h)))"
proof -
have "D.iso (F (ε (G (f ⋆⇩D g) ⋆⇩C G h)) ⋅⇩D η (F (G (f ⋆⇩D g) ⋆⇩C G h)))"
using assms by auto
thus ?thesis
using assms D.inv_comp_right(1) by simp
qed
also have "... = F (G (f ⋆⇩D g) ⋆⇩C G h)"
using assms ηε.triangle_G Foε_ηoF.map_def [of "G (f ⋆⇩D g) ⋆⇩C G h"] by simp
finally show ?thesis
using assms D.comp_cod_arr [of "Φ (G (f ⋆⇩D g), G h)" "F (G (f ⋆⇩D g) ⋆⇩C G h)"]
by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
𝖺⇩D[f, g, h] ⋅⇩D
(((D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h)) ⋅⇩D
(η (f ⋆⇩D g) ⋆⇩D F (G h))) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h))) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "((D.inv (Φ (G (f ⋆⇩D g), G h)) ⋅⇩D Φ (G (f ⋆⇩D g), G h)) ⋅⇩D
(η (f ⋆⇩D g) ⋆⇩D F (G h)))
= η (f ⋆⇩D g) ⋆⇩D F (G h)"
using assms D.comp_inv_arr' Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C
G⇩0_props FF_def D.comp_cod_arr
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(𝖺⇩D[f, g, h] ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D D.inv (η h))) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
proof -
have "((D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h)) ⋅⇩D (η (f ⋆⇩D g) ⋆⇩D F (G h))) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h))
= (D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D D.inv (η h)"
proof -
have "(D.inv (η (f ⋆⇩D g)) ⋆⇩D D.inv (η h)) ⋅⇩D (η (f ⋆⇩D g) ⋆⇩D F (G h))
= (f ⋆⇩D g) ⋆⇩D D.inv (η h)"
using assms D.comp_inv_arr' D.comp_arr_dom
D.interchange [of "D.inv (η (f ⋆⇩D g))" "η (f ⋆⇩D g)" "D.inv (η h)" "F (G h)"]
D.invert_side_of_triangle(2)
by simp
moreover have "((f ⋆⇩D g) ⋆⇩D D.inv (η h)) ⋅⇩D
((D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D F (G h))
= (D.inv (η f) ⋆⇩D D.inv (η g)) ⋆⇩D D.inv (η h)"
using assms D.comp_cod_arr D.comp_arr_dom
D.interchange [of "f ⋆⇩D g" "D.inv (η f) ⋆⇩D D.inv (η g)" "D.inv (η h)" "F (G h)"]
by simp
ultimately show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h)) ⋅⇩D
𝖺⇩D[F (G f), F (G g), F (G h)]) ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
using assms D.assoc_naturality [of "D.inv (η f)" "D.inv (η g)" "D.inv (η h)"]
D.comp_assoc
by simp
finally show ?thesis
using D.comp_assoc by simp
qed
lemma coherence_RHS:
assumes "D.ide f" and "D.ide g" and "D.ide h"
and "src⇩D f = trg⇩D g" and "src⇩D g = trg⇩D h"
shows "F (Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)))
= (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D (D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (Φ (G g, G h))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))"
proof -
have "F (Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)))
= F (G (D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩C
G (D.inv (Φ (G f, G (g ⋆⇩D h)))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h))) ⋅⇩C
(G f ⋆⇩C G (D.inv (η g) ⋆⇩D D.inv (η h)) ⋅⇩C
G (D.inv (Φ (G g, G h))) ⋅⇩C
C.inv (ε (G g ⋆⇩C G h))))"
using assms Φ⇩G_map_simp_ide C.comp_assoc by simp
also have "... = F (G (D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩C
G (D.inv (Φ (G f, G (g ⋆⇩D h)))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h))) ⋅⇩C
(G f ⋆⇩C G (D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩C
(G f ⋆⇩C G (D.inv (Φ (G g, G h)))) ⋅⇩C
(G f ⋆⇩C C.inv (ε (G g ⋆⇩C G h))))"
using assms C.whisker_left by simp
also have "... = F (G (D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h)))) ⋅⇩D
F (G (D.inv (Φ (G f, G (g ⋆⇩D h))))) ⋅⇩D
F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D
F (G f ⋆⇩C G (D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
F (G f ⋆⇩C G (D.inv (Φ (G g, G h)))) ⋅⇩D
F (G f ⋆⇩C C.inv (ε (G g ⋆⇩C G h)))"
proof -
have "C.arr (G (D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩C
G (D.inv (Φ (G f, G (g ⋆⇩D h)))) ⋅⇩C
C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h))) ⋅⇩C
(G f ⋆⇩C G (D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩C
(G f ⋆⇩C G (D.inv (Φ (G g, G h)))) ⋅⇩C
(G f ⋆⇩C C.inv (ε (G g ⋆⇩C G h))))"
using assms G⇩0_props by auto
thus ?thesis
using assms by (metis C.seqE as_nat_trans.preserves_comp_2)
qed
also have "... = F (G (D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h)))) ⋅⇩D
F (G (D.inv (Φ (G f, G (g ⋆⇩D h))))) ⋅⇩D
F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D
(Φ (G f, G (g ⋆⇩D h)) ⋅⇩D
(F (G f) ⋆⇩D F (G (D.inv (η g) ⋆⇩D D.inv (η h)))) ⋅⇩D
((D.inv (Φ (G f, G (F (G g) ⋆⇩D F (G h))))) ⋅⇩D
(Φ (G f, G (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (G (D.inv (Φ (G g, G h)))))) ⋅⇩D
((D.inv (Φ (G f, G (F (G g ⋆⇩C G h))))) ⋅⇩D
(Φ (G f, G (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h)))"
using assms preserves_hcomp G⇩0_props D.comp_assoc by simp
also have "... = F (G (D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h)))) ⋅⇩D
F (G (D.inv (Φ (G f, G (g ⋆⇩D h))))) ⋅⇩D
F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D
(Φ (G f, G (g ⋆⇩D h)) ⋅⇩D
(F (G f) ⋆⇩D F (G (D.inv (η g) ⋆⇩D D.inv (η h)))) ⋅⇩D
(F (G f) ⋆⇩D F (G (D.inv (Φ (G g, G h))))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h)))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h)))"
proof -
have "D.inv (Φ (G f, G (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
Φ (G f, G (F (G g) ⋆⇩D F (G h))) ⋅⇩D
(F (G f) ⋆⇩D F (G (D.inv (Φ (G g, G h)))))
= (F (G f) ⋆⇩D F (G (D.inv (Φ (G g, G h)))))"
proof -
have "D.inv (Φ (G f, G (F (G g) ⋆⇩D F (G h)))) ⋅⇩D Φ (G f, G (F (G g) ⋆⇩D F (G h)))
= F (G f) ⋆⇩D F (G (F (G g) ⋆⇩D F (G h)))"
using assms D.comp_inv_arr D.inv_is_inverse G⇩0_props FF_def by simp
moreover have "... = D.cod (F (G f) ⋆⇩D F (G (D.inv (Φ (G g, G h)))))"
using assms G⇩0_props by simp
moreover have "D.hseq (F (G f)) (F (G (D.inv (Φ (G g, G h)))))"
using assms G⇩0_props by auto
ultimately show ?thesis
using assms D.comp_assoc
D.comp_cod_arr [of "F (G f) ⋆⇩D F (G (D.inv (Φ (G g, G h))))"
"F (G f) ⋆⇩D F (G (F (G g) ⋆⇩D F (G h)))"]
by metis
qed
moreover have "D.inv (Φ (G f, G (F (G g ⋆⇩C G h)))) ⋅⇩D
(Φ (G f, G (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))
= (F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))"
proof -
have "D.inv (Φ (G f, G (F (G g ⋆⇩C G h)))) ⋅⇩D Φ (G f, G (F (G g ⋆⇩C G h)))
= F (G f) ⋆⇩D F (G (F (G g ⋆⇩C G h)))"
using assms D.comp_inv_arr D.inv_is_inverse G⇩0_props by simp
moreover have "... = D.cod (F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))"
using assms by simp
moreover have "D.hseq (F (G f)) (F (C.inv (ε (G g ⋆⇩C G h))))"
using assms by (intro D.hseqI') auto
ultimately show ?thesis
using assms D.comp_assoc
D.comp_cod_arr [of "F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h)))"
"F (G f) ⋆⇩D F (G (F (G g ⋆⇩C G h)))"]
by metis
qed
ultimately show ?thesis by simp
qed
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩D
((D.inv (η (F (G f) ⋆⇩D F (G (g ⋆⇩D h))))) ⋅⇩D
η (F (G f) ⋆⇩D F (G (g ⋆⇩D h)))) ⋅⇩D
D.inv (Φ (G f, G (g ⋆⇩D h)))) ⋅⇩D
((D.inv (η (F (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D
F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h))))) ⋅⇩D
Φ (G f, G (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D
η (g ⋆⇩D h) ⋅⇩D
(D.inv (η g) ⋆⇩D D.inv (η h)) ⋅⇩D
D.inv (η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D
η (F (G g) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G g, G h)) ⋅⇩D
D.inv (η (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h)))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))"
proof -
have "⋀ν. D.arr ν ⟹ F (G ν) = η (D.cod ν) ⋅⇩D ν ⋅⇩D D.inv (η (D.dom ν))"
by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
η_naturality(1-2) η_simps(1,6))
thus ?thesis
using assms D.comp_assoc by simp
qed
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩D
D.inv (Φ (G f, G (g ⋆⇩D h)))) ⋅⇩D
Φ (G f, G (g ⋆⇩D h)) ⋅⇩D
(F (G f) ⋆⇩D
η (g ⋆⇩D h) ⋅⇩D
(D.inv (η g) ⋆⇩D D.inv (η h)) ⋅⇩D
D.inv (η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D
η (F (G g) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G g, G h)) ⋅⇩D
D.inv (η (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h)))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))"
proof -
have "D.inv (η (F (G f) ⋆⇩D F (G (g ⋆⇩D h)))) ⋅⇩D η (F (G f) ⋆⇩D F (G (g ⋆⇩D h))) ⋅⇩D
D.inv (Φ (G f, G (g ⋆⇩D h)))
= D.inv (Φ (G f, G (g ⋆⇩D h)))"
proof -
have "D.inv (η (F (G f) ⋆⇩D F (G (g ⋆⇩D h)))) ⋅⇩D η (F (G f) ⋆⇩D F (G (g ⋆⇩D h))) =
F (G f) ⋆⇩D F (G (g ⋆⇩D h))"
using assms D.comp_inv_arr D.inv_is_inverse by simp
moreover have "... = D.cod (D.inv (Φ (G f, G (g ⋆⇩D h))))"
using assms by simp
ultimately show ?thesis
using assms D.comp_cod_arr [of "D.inv (Φ (G f, G (g ⋆⇩D h)))"]
by (metis D.arr_inv D.comp_assoc D.hcomp_simps(2) D.ideD(1) D.ide_hcomp
G.preserves_ide G_simps(2) G_simps(3) cmp_components_are_iso)
qed
moreover have "(D.inv (η (F (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D
F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h))))) ⋅⇩D
Φ (G f, G (g ⋆⇩D h))
= Φ (G f, G (g ⋆⇩D h))"
proof -
have "D.inv (η (F (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h))))
= F (G f ⋆⇩C G (g ⋆⇩D h))"
proof -
have "D.inv (η (F (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D F (C.inv (ε (G f ⋆⇩C G (g ⋆⇩D h)))) =
D.inv (η (F (G f ⋆⇩C G (g ⋆⇩D h)))) ⋅⇩D D.inv (F (ε (G f ⋆⇩C G (g ⋆⇩D h))))"
using assms by (simp add: preserves_inv)
also have "... = D.inv (F (ε (G f ⋆⇩C G (g ⋆⇩D h))) ⋅⇩D η (F (G f ⋆⇩C G (g ⋆⇩D h))))"
proof -
have "D.iso (F (ε (G f ⋆⇩C G (g ⋆⇩D h))) ⋅⇩D η (F (G f ⋆⇩C G (g ⋆⇩D h))))"
using assms by auto
thus ?thesis
using assms by (simp add: D.inv_comp_right(1))
qed
also have "... = F (G f ⋆⇩C G (g ⋆⇩D h))"
using assms Foε_ηoF.map_def [of "G f ⋆⇩C G (g ⋆⇩D h)"] ηε.triangle_G by simp
finally show ?thesis by blast
qed
moreover have "... = D.cod (Φ (G f, G (g ⋆⇩D h)))"
using assms by simp
moreover have "D.arr (Φ (G f, G (g ⋆⇩D h)))"
using assms by auto
ultimately show ?thesis
using D.comp_cod_arr by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩D
((D.inv (Φ (G f, G (g ⋆⇩D h)))) ⋅⇩D
Φ (G f, G (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D η (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h)) ⋅⇩D
(((F (G f) ⋆⇩D D.inv (η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (Φ (G g, G h)))) ⋅⇩D
((F (G f) ⋆⇩D D.inv (η (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))"
using assms D.whisker_left D.comp_assoc by simp
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(((D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D η (g ⋆⇩D h)))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (Φ (G g, G h))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))"
proof -
have "(D.inv (Φ (G f, G (g ⋆⇩D h))) ⋅⇩D Φ (G f, G (g ⋆⇩D h))) ⋅⇩D (F (G f) ⋆⇩D η (g ⋆⇩D h))
= (F (G f) ⋆⇩D η (g ⋆⇩D h))"
using assms D.comp_inv_arr' G⇩0_props D.comp_cod_arr by simp
moreover have "((F (G f) ⋆⇩D D.inv (η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (Φ (G g, G h)))
= F (G f) ⋆⇩D D.inv (Φ (G g, G h))"
proof -
have "(F (G f) ⋆⇩D D.inv (η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D η (F (G g) ⋆⇩D F (G h)))
= F (G f) ⋆⇩D F (G g) ⋆⇩D F (G h)"
proof -
have "(F (G f) ⋆⇩D D.inv (η (F (G g) ⋆⇩D F (G h)))) ⋅⇩D
(F (G f) ⋆⇩D η (F (G g) ⋆⇩D F (G h)))
= F (G f) ⋆⇩D D.inv (η (F (G g) ⋆⇩D F (G h))) ⋅⇩D η (F (G g) ⋆⇩D F (G h))"
using assms D.interchange [of "F (G f)" "F (G f)"] by simp
also have "... = F (G f) ⋆⇩D F (G g) ⋆⇩D F (G h)"
using assms D.comp_inv_arr' by simp
finally show ?thesis by blast
qed
moreover have "... = D.cod (F (G f) ⋆⇩D D.inv (Φ (G g, G h)))"
using assms by simp
moreover have "D.arr (F (G f) ⋆⇩D D.inv (Φ (G g, G h)))"
using assms by simp
ultimately show ?thesis
using D.comp_cod_arr by simp
qed
moreover have "((F (G f) ⋆⇩D D.inv (η (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))
= D.inv (Φ (G f, G g ⋆⇩C G h))"
proof -
have "(F (G f) ⋆⇩D D.inv (η (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))
= F (G f) ⋆⇩D F (G g ⋆⇩C G h)"
proof -
have "(F (G f) ⋆⇩D D.inv (η (F (G g ⋆⇩C G h)))) ⋅⇩D
(F (G f) ⋆⇩D F (C.inv (ε (G g ⋆⇩C G h))))
= F (G f) ⋆⇩D D.inv (η (F (G g ⋆⇩C G h))) ⋅⇩D F (C.inv (ε (G g ⋆⇩C G h)))"
using assms D.interchange [of "F (G f)" "F (G f)"] by simp
also have "... = F (G f) ⋆⇩D D.inv (η (F (G g ⋆⇩C G h))) ⋅⇩D
D.inv (F (ε (G g ⋆⇩C G h)))"
using assms preserves_inv by simp
also have "... = F (G f) ⋆⇩D D.inv (F (ε (G g ⋆⇩C G h)) ⋅⇩D η (F (G g ⋆⇩C G h)))"
proof -
have "D.iso (F (ε (G g ⋆⇩C G h)) ⋅⇩D η (F (G g ⋆⇩C G h)))"
using assms by auto
thus ?thesis
using assms by (simp add: D.inv_comp_right(1))
qed
also have "... = F (G f) ⋆⇩D F (G g ⋆⇩C G h)"
using assms ηε.triangle_G Foε_ηoF.map_def [of "G g ⋆⇩C G h"] by simp
finally show ?thesis by blast
qed
thus ?thesis using assms D.comp_cod_arr by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (Φ (G g, G h))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))"
proof -
have "((D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D η (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))
= D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h)"
using assms D.comp_cod_arr D.comp_arr_dom D.interchange
proof -
have "((D.inv (η f) ⋆⇩D D.inv (η (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D η (g ⋆⇩D h))) ⋅⇩D
(F (G f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))
= (D.inv (η f) ⋅⇩D F (G f) ⋆⇩D D.inv (η (g ⋆⇩D h)) ⋅⇩D η (g ⋆⇩D h)) ⋅⇩D
(F (G f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))"
using assms D.interchange by simp
also have "... = (D.inv (η f) ⋆⇩D g ⋆⇩D h) ⋅⇩D (F (G f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))"
using assms D.comp_inv_arr' D.comp_arr_dom by simp
also have "... = D.inv (η f) ⋅⇩D F (G f) ⋆⇩D (g ⋆⇩D h) ⋅⇩D (D.inv (η g) ⋆⇩D D.inv (η h))"
using assms D.interchange by simp
also have "... = D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h)"
using assms D.comp_arr_dom D.comp_cod_arr by simp
finally show ?thesis by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
finally show ?thesis by blast
qed
interpretation G: pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C G Φ⇩G.map
proof
show "⋀f g h. ⟦D.ide f; D.ide g; D.ide h; src⇩D f = trg⇩D g; src⇩D g = trg⇩D h⟧
⟹ G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C (Φ⇩G.map (f, g) ⋆⇩C G h) =
Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)) ⋅⇩C 𝖺⇩C[G f, G g, G h]"
proof -
fix f g h
assume f: "D.ide f" and g: "D.ide g" and h: "D.ide h"
assume fg: "src⇩D f = trg⇩D g" and gh: "src⇩D g = trg⇩D h"
have "F (G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C (Φ⇩G.map (f, g) ⋆⇩C G h)) =
F (Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)) ⋅⇩C 𝖺⇩C[G f, G g, G h])"
proof -
have "F (G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C (Φ⇩G.map (f, g) ⋆⇩C G h))
= (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
𝖺⇩D[F (G f), F (G g), F (G h)] ⋅⇩D
(D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))"
using f g h fg gh coherence_LHS by simp
also have "... = (η (f ⋆⇩D g ⋆⇩D h) ⋅⇩D
(D.inv (η f) ⋆⇩D D.inv (η g) ⋆⇩D D.inv (η h))) ⋅⇩D
((F (G f) ⋆⇩D D.inv (Φ (G g, G h))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))) ⋅⇩D
F 𝖺⇩C[G f, G g, G h]"
proof -
have "𝖺⇩D[F (G f), F (G g), F (G h)] ⋅⇩D (D.inv (Φ (G f, G g)) ⋆⇩D F (G h)) ⋅⇩D
D.inv (Φ (G f ⋆⇩C G g, G h))
= 𝖺⇩D[F (G f), F (G g), F (G h)] ⋅⇩D D.inv (Φ (G f ⋆⇩C G g, G h) ⋅⇩D
(Φ (G f, G g) ⋆⇩D F (G h)))"
proof -
have "D.iso (Φ (G f ⋆⇩C G g, G h) ⋅⇩D (Φ (G f, G g) ⋆⇩D F (G h)))"
using f g h fg gh Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩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⇩S⇩b⇩C D.inv_comp_right(1))
qed
also have "... = D.inv (Φ (G f, G g ⋆⇩C G h) ⋅⇩D (F (G f) ⋆⇩D Φ (G g, G h))) ⋅⇩D
F 𝖺⇩C[G f, G g, G h]"
using f g h fg gh cmp_simps assoc_coherence D.comp_assoc D.isos_compose
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C FF_def
D.invert_opposite_sides_of_square
by simp
also have "... = ((F (G f) ⋆⇩D D.inv (Φ (G g, G h))) ⋅⇩D
D.inv (Φ (G f, G g ⋆⇩C G h))) ⋅⇩D
F 𝖺⇩C[G f, G g, G h]"
proof -
have "D.iso (Φ (G f, G g ⋆⇩C G h) ⋅⇩D (F (G f) ⋆⇩D Φ (G g, G h)))"
using f g h fg gh Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C FF_def
by auto
thus ?thesis
using f g h fg gh Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C FF_def
by (simp add: C.VV.arr_char⇩S⇩b⇩C D.comp_assoc D.inv_comp_right(1))
qed
finally show ?thesis
using D.comp_assoc by simp
qed
also have "... = F (Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h))) ⋅⇩D
F 𝖺⇩C[G f, G g, G h]"
using f g h fg gh coherence_RHS D.comp_assoc by simp
also have "... = F ((Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h))) ⋅⇩C
𝖺⇩C[G f, G g, G h])"
using f g h fg gh Φ⇩G_simps by auto
also have "... = F (Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)) ⋅⇩C
𝖺⇩C[G f, G g, G h])"
using C.comp_assoc by simp
finally show ?thesis by simp
qed
moreover have "C.par (G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C
(Φ⇩G.map (f, g) ⋆⇩C G h))
(Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)) ⋅⇩C 𝖺⇩C[G f, G g, G h])"
using f g h fg gh Φ⇩G_simps by auto
ultimately show "G 𝖺⇩D[f, g, h] ⋅⇩C Φ⇩G.map (f ⋆⇩D g, h) ⋅⇩C (Φ⇩G.map (f, g) ⋆⇩C G h) =
Φ⇩G.map (f, g ⋆⇩D h) ⋅⇩C (G f ⋆⇩C Φ⇩G.map (g, h)) ⋅⇩C 𝖺⇩C[G f, G g, G h]"
using is_faithful by blast
qed
qed
interpretation GF: composite_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ G Φ⇩G.map
..
interpretation FG: composite_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G.map F Φ
..
interpretation I⇩C: identity_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C ..
interpretation I⇩D: identity_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D ..
lemma η_equals_FG_unit:
assumes "D.obj a"
shows "η a = FG.unit a"
proof (intro FG.unit_eqI)
show "D.obj a" by fact
show "D.iso (η a)"
using assms by auto
show "«η a : FG.map⇩0 a ⇒⇩D FG.map a»"
using assms η_in_hom [of a] FG.map⇩0_def G⇩0_props D.obj_def
by (metis D.ideD(2) D.ideD(3) D.objE D.vconn_implies_hpar(3) η.preserves_cod
η_simps(5))
show "η a ⋅⇩D 𝗂⇩D[FG.map⇩0 a] = (FG.map 𝗂⇩D[a] ⋅⇩D FG.cmp (a, a)) ⋅⇩D (η a ⋆⇩D η a)"
proof -
have "(FG.map 𝗂⇩D[a] ⋅⇩D FG.cmp (a, a)) ⋅⇩D (η a ⋆⇩D η a) =
F (G 𝗂⇩D[a]) ⋅⇩D FG.cmp (a, a) ⋅⇩D (η a ⋆⇩D η a)"
using assms D.comp_assoc by simp
also have "... = F (G 𝗂⇩D[a]) ⋅⇩D F (G (I⇩D.cmp (a, a))) ⋅⇩D F (Φ⇩G.map (a, a)) ⋅⇩D
Φ (G a, G a) ⋅⇩D (η a ⋆⇩D η a)"
using assms FG.cmp_def D.comp_assoc D.VV.arr_char⇩S⇩b⇩C D.VV.dom_char⇩S⇩b⇩C D.VV.cod_char⇩S⇩b⇩C
by auto
also have "... = F (G 𝗂⇩D[a]) ⋅⇩D F (G (I⇩D.cmp (a, a))) ⋅⇩D η (a ⋆⇩D a)"
using assms η_hcomp by auto
also have "... = F (G 𝗂⇩D[a]) ⋅⇩D η (a ⋆⇩D a)"
using assms D.comp_cod_arr by auto
also have "... = η a ⋅⇩D 𝗂⇩D[a]"
using assms η_naturality [of "𝗂⇩D[a]"] by simp
also have "... = η a ⋅⇩D 𝗂⇩D[FG.map⇩0 a]"
using assms ‹«η a : FG.map⇩0 a ⇒⇩D FG.map a»› by fastforce
finally show ?thesis by simp
qed
qed
lemma ε_hcomp':
assumes "C.ide g" and "C.ide f" and "src⇩C g = trg⇩C f"
shows "ε (g ⋆⇩C f) ⋅⇩C GF.cmp (g, f) = ε g ⋆⇩C ε f"
proof -
have "ε (g ⋆⇩C f) ⋅⇩C GF.cmp (g, f)
= (ε g ⋆⇩C ε f) ⋅⇩C C.inv (Φ⇩G.map (F g, F f)) ⋅⇩C C.inv (G (Φ (g, f))) ⋅⇩C
G (F (g ⋆⇩C f)) ⋅⇩C G (Φ (g, f)) ⋅⇩C Φ⇩G.map (F g, F f)"
using assms ε_hcomp GF.cmp_def C.VV.arr_char⇩S⇩b⇩C C.comp_cod_arr
C.comp_inv_arr' C.comp_assoc C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
by simp
also have "... = (ε g ⋆⇩C ε f) ⋅⇩C C.inv (Φ⇩G.map (F g, F f)) ⋅⇩C (C.inv (G (Φ (g, f))) ⋅⇩C
G (F (g ⋆⇩C f)) ⋅⇩C G (Φ (g, f))) ⋅⇩C Φ⇩G.map (F g, F f)"
using C.comp_assoc by simp
also have "... = (ε g ⋆⇩C ε f) ⋅⇩C C.inv (Φ⇩G.map (F g, F f)) ⋅⇩C (C.inv (G (Φ (g, f))) ⋅⇩C
G (Φ (g, f))) ⋅⇩C Φ⇩G.map (F g, F f)"
using assms C.comp_ide_arr [of "G (F (g ⋆⇩C f))" "G (Φ (g, f))"]
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
by simp
also have "... = (ε g ⋆⇩C ε f) ⋅⇩C C.inv (Φ⇩G.map (F g, F f)) ⋅⇩C Φ⇩G.map (F g, F f)"
proof -
have "C.inv (G (Φ (g, f))) ⋅⇩C G (Φ (g, f)) = G (F g ⋆⇩D F f)"
using assms C.comp_inv_arr' cmp_components_are_iso C.inv_is_inverse
C.VV.arr_char⇩S⇩b⇩C C.VV.ide_char⇩S⇩b⇩C cmp_simps(4)
by auto
moreover have "... = C.cod (Φ⇩G.map (F g, F f))"
using assms by simp
ultimately have "(C.inv (G (Φ (g, f))) ⋅⇩C G (Φ (g, f))) ⋅⇩C Φ⇩G.map (F g, F f) =
Φ⇩G.map (F g, F f)"
using assms C.comp_cod_arr [of "Φ⇩G.map (F g, F f)" "G (F g ⋆⇩D F f)"]
C.ideD(1) G.cmp_simps(1) preserves_ide preserves_src preserves_trg
by presburger
thus ?thesis by simp
qed
also have "... = ε g ⋆⇩C ε f"
using assms C.comp_inv_arr' C.comp_arr_dom [of "ε g ⋆⇩C ε f" "G (F g) ⋆⇩C G (F f)"]
by simp
finally show ?thesis by simp
qed
lemma ε_inverts_GF_unit:
assumes "C.obj a"
shows "ε a ⋅⇩C GF.unit a = a"
proof -
have "ε a ⋅⇩C GF.unit a = I⇩C.unit a"
proof (intro I⇩C.unit_eqI)
show "C.obj a" by fact
show 1: "«ε a ⋅⇩C GF.unit a : I⇩C.map⇩0 a ⇒⇩C I⇩C.map a»"
proof -
have "src⇩C (G (F a)) = src⇩C (I⇩C.map a)"
using assms G⇩0_props C.obj_def' by simp
thus ?thesis
using assms I⇩C.map⇩0_def GF.map⇩0_def GF.unit_in_hom
by (intro C.comp_in_homI') auto
qed
show "C.iso (ε a ⋅⇩C GF.unit a)"
using assms ε_in_hom GF.unit_char(2)
by (intro C.isos_compose) auto
show "(ε a ⋅⇩C GF.unit a) ⋅⇩C 𝗂⇩C[I⇩C.map⇩0 a]
= (I⇩C.map 𝗂⇩C[a] ⋅⇩C I⇩C.cmp (a, a)) ⋅⇩C (ε a ⋅⇩C GF.unit a ⋆⇩C ε a ⋅⇩C GF.unit a)"
proof -
have "(I⇩C.map 𝗂⇩C[a] ⋅⇩C I⇩C.cmp (a, a)) ⋅⇩C (ε a ⋅⇩C GF.unit a ⋆⇩C ε a ⋅⇩C GF.unit a) =
𝗂⇩C[a] ⋅⇩C (a ⋆⇩C a) ⋅⇩C (ε a ⋅⇩C GF.unit a ⋆⇩C ε a ⋅⇩C GF.unit a)"
using assms C.comp_assoc by simp
also have "... = 𝗂⇩C[a] ⋅⇩C (ε a ⋅⇩C GF.unit a ⋆⇩C ε a ⋅⇩C GF.unit a)"
proof -
have "C.hseq (ε a ⋅⇩C GF.unit a) (ε a ⋅⇩C GF.unit a)"
using assms GF.unit_simps C.iso_is_arr ‹C.iso (ε a ⋅⇩C GF.unit a)›
by (intro C.hseqI') auto
moreover have "a ⋆⇩C a = C.cod (ε a ⋅⇩C GF.unit a ⋆⇩C ε a ⋅⇩C GF.unit a)"
proof -
have "C.cod (ε a ⋅⇩C GF.unit a) = a"
using assms 1 C.obj_simps by auto
moreover have "C.hseq (ε a ⋅⇩C GF.unit a) (ε a ⋅⇩C GF.unit a)"
using assms 1 C.src_dom [of "ε a ⋅⇩C GF.unit a"] C.trg_dom [of "ε a ⋅⇩C GF.unit a"]
apply (intro C.hseqI')
by auto force
ultimately show ?thesis by auto
qed
ultimately show ?thesis
using C.comp_cod_arr by simp
qed
also have "... = 𝗂⇩C[a] ⋅⇩C (ε a ⋆⇩C ε a) ⋅⇩C (GF.unit a ⋆⇩C GF.unit a)"
using assms C.interchange [of "ε a" "GF.unit a" "ε a" "GF.unit a"]
by (simp add: C.iso_is_arr ‹C.iso (ε a ⋅⇩C GF.unit a)›)
also have "... = 𝗂⇩C[a] ⋅⇩C (ε (a ⋆⇩C a) ⋅⇩C GF.cmp (a, a)) ⋅⇩C (GF.unit a ⋆⇩C GF.unit a)"
using assms ε_hcomp' [of a a] by auto
also have "... = (𝗂⇩C[a] ⋅⇩C ε (a ⋆⇩C a)) ⋅⇩C GF.cmp (a, a) ⋅⇩C (GF.unit a ⋆⇩C GF.unit a)"
using C.comp_assoc by simp
also have "... = (ε a ⋅⇩C G (F 𝗂⇩C[a])) ⋅⇩C GF.cmp (a, a) ⋅⇩C (GF.unit a ⋆⇩C GF.unit a)"
using assms ε_naturality [of "𝗂⇩C[a]"] by simp
also have "... = ε a ⋅⇩C (G (F 𝗂⇩C[a]) ⋅⇩C GF.cmp (a, a)) ⋅⇩C (GF.unit a ⋆⇩C GF.unit a)"
using C.comp_assoc by simp
also have "... = ε a ⋅⇩C GF.unit a ⋅⇩C 𝗂⇩C[GF.map⇩0 a]"
using assms GF.unit_char [of a] by simp
also have "... = (ε a ⋅⇩C GF.unit a) ⋅⇩C 𝗂⇩C[I⇩C.map⇩0 a]"
proof -
have "GF.map⇩0 a = I⇩C.map⇩0 a"
using assms G⇩0_props(2) [of a] GF.map⇩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⇩C.unit_char' by simp
qed
lemma η_respects_comp:
assumes "D.ide f" and "D.ide g" and "src⇩D g = trg⇩D f"
shows "(𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D η (g ⋆⇩D f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]) ⋅⇩D ((g ⋆⇩D f) ⋆⇩D src⇩D f)
= (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
and "(trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D 𝖺⇩D[trg⇩D g, F (G g), F (G f)] ⋅⇩D
(𝗅⇩D⇧-⇧1[F (G g)] ⋅⇩D η g ⋅⇩D 𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D D.inv 𝖺⇩D[g, src⇩D g, F (G f)] ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)] ⋅⇩D η f ⋅⇩D 𝗋⇩D[f]) ⋅⇩D 𝖺⇩D[g, f, src⇩D f]
= (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
proof -
show "(𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D η (g ⋆⇩D f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]) ⋅⇩D ((g ⋆⇩D f) ⋆⇩D src⇩D f)
= (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
proof -
have "(𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D η (g ⋆⇩D f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]) ⋅⇩D ((g ⋆⇩D f) ⋆⇩D src⇩D f)
= 𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D η (g ⋆⇩D f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
using assms D.comp_assoc D.comp_arr_dom by simp
also have 1: "... = (𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D F (Φ⇩G.map (g, f))) ⋅⇩D
Φ (G g, G f) ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
using assms η_hcomp D.comp_assoc by simp
also have "... = (trg⇩D g ⋆⇩D F (Φ⇩G.map (g, f))) ⋅⇩D (𝗅⇩D⇧-⇧1[F (G g ⋆⇩C G f)] ⋅⇩D
Φ (G g, G f)) ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
proof -
have "𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D F (Φ⇩G.map (g, f)) =
(trg⇩D g ⋆⇩D F (Φ⇩G.map (g, f))) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G g ⋆⇩C G f)]"
using assms G⇩0_props D.lunit'_naturality [of "F (Φ⇩G.map (g, f))"] Φ⇩G_in_hom [of g f]
by auto
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((trg⇩D g ⋆⇩D F (Φ⇩G.map (g, f))) ⋅⇩D (trg⇩D g ⋆⇩D Φ (G g, G f))) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
proof -
have "𝗅⇩D⇧-⇧1[F (G g ⋆⇩C G f)] ⋅⇩D Φ (G g, G f)
= (trg⇩D g ⋆⇩D Φ (G g, G f)) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)]"
using assms D.lunit'_naturality [of "Φ (G g, G f)"] G⇩0_props by fastforce
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg⇩D g ⋆⇩D F (Φ⇩G.map (g, f)) ⋅⇩D Φ (G g, G f)) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
using assms 1 D.whisker_left [of "trg⇩D g" "F (Φ⇩G.map (g, f))" "Φ (G g, G f)"]
by force
also have "... = (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D (η g ⋆⇩D η f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
proof -
have 1: "FG.cmp (g, f) = (F (G (g ⋆⇩D f)) ⋅⇩D F (Φ⇩G.map (g, f))) ⋅⇩D Φ (G g, G f)"
using assms FG.cmp_def D.VV.arr_char⇩S⇩b⇩C D.VV.dom_char⇩S⇩b⇩C D.comp_assoc by simp
also have "... = F (Φ⇩G.map (g, f)) ⋅⇩D Φ (G g, G f)"
proof -
have "D.cod (F (Φ⇩G (g, f))) = F (G (g ⋆⇩D f))"
using assms 1
by (metis (mono_tags, lifting) D.cod_eqI D.ideD(1) D.ide_hcomp D.seqE
FG.cmp_simps'(1) G.preserves_ide preserves_ide)
thus ?thesis
using assms D.comp_cod_arr [of "F (Φ⇩G.map (g, f))" "F (G (g ⋆⇩D f))"]
by fastforce
qed
finally have "FG.cmp (g, f) = F (Φ⇩G.map (g, f)) ⋅⇩D Φ (G g, G f)" by simp
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
show "(trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝖺⇩D[trg⇩D g, F (G g), F (G f)] ⋅⇩D
(𝗅⇩D⇧-⇧1[F (G g)] ⋅⇩D η g ⋅⇩D 𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D
D.inv 𝖺⇩D[g, src⇩D g, F (G f)] ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)] ⋅⇩D η f ⋅⇩D 𝗋⇩D[f]) ⋅⇩D
𝖺⇩D[g, f, src⇩D f]
= (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D
(η g ⋆⇩D η f) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
proof -
have "(trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝖺⇩D[trg⇩D g, F (G g), F (G f)] ⋅⇩D
(𝗅⇩D⇧-⇧1[F (G g)] ⋅⇩D η g ⋅⇩D 𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D
D.inv 𝖺⇩D[g, src⇩D g, F (G f)] ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)] ⋅⇩D η f ⋅⇩D 𝗋⇩D[f]) ⋅⇩D
𝖺⇩D[g, f, src⇩D f]
= (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝖺⇩D[trg⇩D g, F (G g), F (G f)] ⋅⇩D
(𝗅⇩D⇧-⇧1[F (G g)] ⋆⇩D F (G f)) ⋅⇩D
(η g ⋆⇩D F (G f)) ⋅⇩D
(𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D
D.inv 𝖺⇩D[g, src⇩D g, F (G f)] ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)]) ⋅⇩D
(g ⋆⇩D η f) ⋅⇩D
(g ⋆⇩D 𝗋⇩D[f]) ⋅⇩D
𝖺⇩D[g, f, src⇩D f]"
using assms D.comp_assoc D.whisker_right D.whisker_left by simp
also have "... = (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
(𝖺⇩D[trg⇩D g, F (G g), F (G f)] ⋅⇩D
(𝗅⇩D⇧-⇧1[F (G g)] ⋆⇩D F (G f))) ⋅⇩D
(η g ⋆⇩D F (G f)) ⋅⇩D
(𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D
D.inv 𝖺⇩D[g, src⇩D g, F (G f)] ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)]) ⋅⇩D
(g ⋆⇩D η f) ⋅⇩D
((g ⋆⇩D 𝗋⇩D[f]) ⋅⇩D
𝖺⇩D[g, f, src⇩D f])"
using assms D.comp_cod_arr D.VV.ide_char⇩S⇩b⇩C D.VV.arr_char⇩S⇩b⇩C D.VV.dom_char⇩S⇩b⇩C
FG.FF_def G⇩0_props D.comp_assoc
by presburger
also have "... = (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D
(η g ⋆⇩D F (G f)) ⋅⇩D
((𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D
D.inv 𝖺⇩D[g, src⇩D g, F (G f)]) ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)]) ⋅⇩D
(g ⋆⇩D η f) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
proof -
have "(g ⋆⇩D 𝗋⇩D[f]) ⋅⇩D 𝖺⇩D[g, f, src⇩D f] = 𝗋⇩D[g ⋆⇩D f]"
using assms D.runit_hcomp by simp
moreover have "𝖺⇩D[trg⇩D g, F (G g), F (G f)] ⋅⇩D (𝗅⇩D⇧-⇧1[F (G g)] ⋆⇩D F (G f)) =
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)]"
using assms D.lunit_hcomp [of "F (G g)" "F (G f)"] G⇩0_props by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D
(η g ⋆⇩D F (G f)) ⋅⇩D
(((g ⋆⇩D 𝗅⇩D[F (G f)]) ⋅⇩D
(g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)])) ⋅⇩D
(g ⋆⇩D η f)) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
proof -
have "(𝗋⇩D[g] ⋆⇩D F (G f)) ⋅⇩D D.inv 𝖺⇩D[g, src⇩D g, F (G f)] = g ⋆⇩D 𝗅⇩D[F (G f)]"
using assms D.triangle' [of g "F (G f)"] G⇩0_props by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D
((η g ⋆⇩D F (G f)) ⋅⇩D
(g ⋆⇩D η f)) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]"
proof -
have "((g ⋆⇩D 𝗅⇩D[F (G f)]) ⋅⇩D (g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)])) ⋅⇩D (g ⋆⇩D η f) = g ⋆⇩D η f"
using assms D.interchange [of g g "𝗅⇩D[F (G f)]" "𝗅⇩D⇧-⇧1[F (G f)]"]
D.comp_arr_inv' D.comp_cod_arr
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G g) ⋆⇩D F (G f)] ⋅⇩D
(η g ⋆⇩D η f) ⋅⇩D 𝗋⇩D[g ⋆⇩D f]"
using assms D.interchange [of "η g" g "F (G f)" "η f"] D.comp_arr_dom D.comp_cod_arr
by simp
finally show ?thesis by simp
qed
qed
lemma η_respects_unit:
assumes "D.obj a"
shows "(a ⋆⇩D FG.unit a) ⋅⇩D 𝗋⇩D⇧-⇧1[a] ⋅⇩D 𝗅⇩D[a] =
(𝗅⇩D⇧-⇧1[FG.map (D.cod a)] ⋅⇩D η a ⋅⇩D 𝗋⇩D[D.dom a]) ⋅⇩D (I⇩D.unit a ⋆⇩D a)"
proof -
have "(𝗅⇩D⇧-⇧1[FG.map (D.cod a)] ⋅⇩D η a ⋅⇩D 𝗋⇩D[D.dom a]) ⋅⇩D (I⇩D.unit a ⋆⇩D a) =
(𝗅⇩D⇧-⇧1[F (G a)] ⋅⇩D η a) ⋅⇩D 𝗋⇩D[a]"
using assms I⇩D.lunit_coherence I⇩D.unit_char' D.comp_arr_dom D.comp_assoc by auto
also have "... = ((a ⋆⇩D η a) ⋅⇩D 𝗅⇩D⇧-⇧1[a]) ⋅⇩D 𝗋⇩D[a]"
using assms D.lunit'_naturality [of "η a"] by auto
also have "... = (a ⋆⇩D η a) ⋅⇩D 𝗅⇩D⇧-⇧1[a] ⋅⇩D 𝗋⇩D[a]"
using D.comp_assoc by simp
also have "... = (a ⋆⇩D η a) ⋅⇩D 𝗋⇩D⇧-⇧1[a] ⋅⇩D 𝗅⇩D[a]"
using assms D.unitor_coincidence by simp
also have "... = (a ⋆⇩D FG.unit a) ⋅⇩D 𝗋⇩D⇧-⇧1[a] ⋅⇩D 𝗅⇩D[a]"
using assms η_equals_FG_unit by simp
finally show ?thesis by simp
qed
lemma ε_respects_comp:
assumes "C.ide f" and "C.ide g" and "src⇩C g = trg⇩C f"
shows "(trg⇩C g ⋆⇩C g ⋆⇩C f) ⋅⇩C 𝖺⇩C[trg⇩C g, g, f] ⋅⇩C (𝗅⇩C⇧-⇧1[g] ⋅⇩C ε g ⋅⇩C 𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f] ⋅⇩C (G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[G (F g), G (F f), src⇩C f]
= 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C (ε g ⋆⇩C ε f) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
and "(𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C ε (g ⋆⇩C f) ⋅⇩C 𝗋⇩C[G (F (g ⋆⇩C f))]) ⋅⇩C (GF.cmp (g, f) ⋆⇩C src⇩C f)
= 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C (ε g ⋆⇩C ε f) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
proof -
have "(trg⇩C g ⋆⇩C g ⋆⇩C f) ⋅⇩C
𝖺⇩C[trg⇩C g, g, f] ⋅⇩C
(𝗅⇩C⇧-⇧1[g] ⋅⇩C ε g ⋅⇩C 𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f] ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[G (F g), G (F f), src⇩C f]
= ((trg⇩C g ⋆⇩C g ⋆⇩C f) ⋅⇩C
𝖺⇩C[trg⇩C g, g, f]) ⋅⇩C
(𝗅⇩C⇧-⇧1[g] ⋅⇩C ε g ⋅⇩C 𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f] ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[G (F g), G (F f), src⇩C f]"
using assms C.comp_assoc by simp
also have "... = 𝖺⇩C[trg⇩C g, g, f] ⋅⇩C
(𝗅⇩C⇧-⇧1[g] ⋅⇩C ε g ⋅⇩C 𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f] ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[G (F g), G (F f), src⇩C f]"
using assms C.comp_cod_arr by simp
also have "... = (𝖺⇩C[trg⇩C g, g, f] ⋅⇩C
(𝗅⇩C⇧-⇧1[g] ⋆⇩C f)) ⋅⇩C
(ε g ⋆⇩C f) ⋅⇩C
(𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f] ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f]) ⋅⇩C
(G (F g) ⋆⇩C ε f) ⋅⇩C
((G (F g) ⋆⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[G (F g), G (F f), src⇩C f])"
using assms C.whisker_left C.whisker_right C.comp_assoc by simp
also have "... = 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C
(ε g ⋆⇩C f) ⋅⇩C
((𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f]) ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f]) ⋅⇩C
(G (F g) ⋆⇩C ε f) ⋅⇩C
𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
using assms G⇩0_props C.lunit_hcomp C.runit_hcomp C.comp_assoc by simp
also have "... = 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C
(ε g ⋆⇩C f) ⋅⇩C
(((G (F g) ⋆⇩C 𝗅⇩C[f]) ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f])) ⋅⇩C
(G (F g) ⋆⇩C ε f)) ⋅⇩C
𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
using assms G⇩0_props C.triangle' C.comp_assoc by simp
also have "... = 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C
((ε g ⋆⇩C f) ⋅⇩C
(G (F g) ⋆⇩C ε f)) ⋅⇩C
𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
proof -
have "(G (F g) ⋆⇩C 𝗅⇩C[f]) ⋅⇩C (G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f]) = G (F g) ⋆⇩C f"
using assms C.whisker_left [of "G (F g)" "𝗅⇩C[f]" "𝗅⇩C⇧-⇧1[f]"] C.comp_arr_inv'
by simp
moreover have "... = C.cod (G (F g) ⋆⇩C ε f)"
using assms G⇩0_props by auto
moreover have "C.hseq (G (F g)) (ε f)"
using assms G⇩0_props by simp
ultimately have "((G (F g) ⋆⇩C 𝗅⇩C[f]) ⋅⇩C (G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f])) ⋅⇩C (G (F g) ⋆⇩C ε f) =
(G (F g) ⋆⇩C ε f)"
using assms G⇩0_props C.comp_cod_arr by presburger
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C (ε g ⋆⇩C ε f) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
using assms C.interchange [of "ε g" "G (F g)" f "ε f"] C.comp_cod_arr C.comp_arr_dom
by simp
finally show "(trg⇩C g ⋆⇩C g ⋆⇩C f) ⋅⇩C
𝖺⇩C[trg⇩C g, g, f] ⋅⇩C
(𝗅⇩C⇧-⇧1[g] ⋅⇩C ε g ⋅⇩C 𝗋⇩C[G (F g)] ⋆⇩C f) ⋅⇩C
C.inv 𝖺⇩C[G (F g), src⇩C g, f] ⋅⇩C
(G (F g) ⋆⇩C 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[G (F g), G (F f), src⇩C f]
= 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C (ε g ⋆⇩C ε f) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
by simp
show "(𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C ε (g ⋆⇩C f) ⋅⇩C 𝗋⇩C[G (F (g ⋆⇩C f))]) ⋅⇩C (GF.cmp (g, f) ⋆⇩C src⇩C f)
= ..."
proof -
have "(𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C ε (g ⋆⇩C f) ⋅⇩C 𝗋⇩C[G (F (g ⋆⇩C f))]) ⋅⇩C (GF.cmp (g, f) ⋆⇩C src⇩C f) =
𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C ε (g ⋆⇩C f) ⋅⇩C 𝗋⇩C[G (F (g ⋆⇩C f))] ⋅⇩C (GF.cmp (g, f) ⋆⇩C src⇩C f)"
using assms C.comp_assoc by simp
also have "... = 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C (ε (g ⋆⇩C f) ⋅⇩C GF.cmp (g, f)) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
proof -
have "src⇩C (GF.cmp (g, f)) = src⇩C f"
using assms G⇩0_props by simp
hence "𝗋⇩C[G (F (g ⋆⇩C f))] ⋅⇩C (GF.cmp (g, f) ⋆⇩C src⇩C f) =
GF.cmp (g, f) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
using assms C.runit_naturality [of "GF.cmp (g, f)"] C.VV.arr_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
GF.cmp_simps'(1,4-5)
by simp
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = 𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C (ε g ⋆⇩C ε f) ⋅⇩C 𝗋⇩C[G (F g) ⋆⇩C G (F f)]"
using assms ε_hcomp' by simp
ultimately show ?thesis by simp
qed
qed
lemma ε_respects_unit:
assumes "C.obj a"
shows "(a ⋆⇩C I⇩C.unit a) ⋅⇩C 𝗋⇩C⇧-⇧1[a] ⋅⇩C 𝗅⇩C[a] =
(𝗅⇩C⇧-⇧1[C.cod a] ⋅⇩C ε a ⋅⇩C 𝗋⇩C[GF.map (C.dom a)]) ⋅⇩C (GF.unit a ⋆⇩C a)"
proof -
have "(𝗅⇩C⇧-⇧1[C.cod a] ⋅⇩C ε a ⋅⇩C 𝗋⇩C[GF.map (C.dom a)]) ⋅⇩C (GF.unit a ⋆⇩C a) =
𝗅⇩C⇧-⇧1[a] ⋅⇩C ε a ⋅⇩C 𝗋⇩C[G (F a)] ⋅⇩C (GF.unit a ⋆⇩C a)"
using assms C.comp_assoc by auto
also have "... = (𝗅⇩C⇧-⇧1[a] ⋅⇩C ε a) ⋅⇩C GF.unit a ⋅⇩C 𝗋⇩C[a]"
proof -
have "src⇩C (GF.unit a) = a"
using assms GF.unit_simps(2) GF.map⇩0_def [of a] G⇩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 ⋆⇩C ε a) ⋅⇩C (𝗅⇩C⇧-⇧1[G (F a)] ⋅⇩C GF.unit a) ⋅⇩C 𝗋⇩C[a]"
proof -
have "𝗅⇩C⇧-⇧1[a] ⋅⇩C ε a = (a ⋆⇩C ε a) ⋅⇩C 𝗅⇩C⇧-⇧1[G (F a)]"
using assms C.lunit'_naturality [of "ε a"] by auto
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = ((a ⋆⇩C ε a) ⋅⇩C (a ⋆⇩C GF.unit a)) ⋅⇩C 𝗅⇩C⇧-⇧1[a] ⋅⇩C 𝗋⇩C[a]"
proof -
have "𝗅⇩C⇧-⇧1[G (F a)] ⋅⇩C GF.unit a = (a ⋆⇩C GF.unit a) ⋅⇩C 𝗅⇩C⇧-⇧1[a]"
using assms C.lunit'_naturality [of "GF.unit a"] G⇩0_props C.obj_simps
by (simp add: GF.map⇩0_def)
thus ?thesis
using C.comp_assoc by simp
qed
also have "... = (a ⋆⇩C ε a ⋅⇩C GF.unit a) ⋅⇩C 𝗅⇩C⇧-⇧1[a] ⋅⇩C 𝗋⇩C[a]"
using assms C.interchange [of a a "ε a" "GF.unit a"] by force
also have "... = (a ⋆⇩C ε a ⋅⇩C GF.unit a) ⋅⇩C 𝗋⇩C⇧-⇧1[a] ⋅⇩C 𝗅⇩C[a]"
using assms C.unitor_coincidence by simp
also have "... = (a ⋆⇩C I⇩C.unit a) ⋅⇩C 𝗋⇩C⇧-⇧1[a] ⋅⇩C 𝗅⇩C[a]"
using assms ε_inverts_GF_unit I⇩C.unit_char' by simp
finally show ?thesis by simp
qed
abbreviation counit⇩0'
where "counit⇩0' ≡ λb. b"
abbreviation counit⇩1'
where "counit⇩1' ≡ λg. 𝗅⇩D⇧-⇧1[F (G g)] ⋅⇩D η g ⋅⇩D 𝗋⇩D[g]"
interpretation ε: pseudonatural_equivalence
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
FG.map FG.cmp I⇩D.map I⇩D.cmp counit⇩0' counit⇩1'
proof
show "⋀a. D.obj a ⟹ D.ide a"
by auto
show "⋀f. D.ide f ⟹ D.iso (𝗅⇩D⇧-⇧1[F (G f)] ⋅⇩D η f ⋅⇩D 𝗋⇩D[f])"
using D.iso_lunit' D.iso_runit η_in_hom
by (intro D.isos_compose D.seqI) auto
show "⋀a. D.obj a ⟹ «a : src⇩D (FG.map a) →⇩D src⇩D (I⇩D.map a)»"
using D.obj_def G⇩0_props(1) by auto
show "⋀f. D.ide f ⟹ «𝗅⇩D⇧-⇧1[F (G f)] ⋅⇩D η f ⋅⇩D 𝗋⇩D[f] :
I⇩D.map f ⋆⇩D src⇩D f ⇒⇩D trg⇩D f ⋆⇩D FG.map f»"
using G⇩0_props
by (intro D.comp_in_homI') auto
show "⋀μ. D.arr μ ⟹
(𝗅⇩D⇧-⇧1[F (G (D.cod μ))] ⋅⇩D η (D.cod μ) ⋅⇩D 𝗋⇩D[D.cod μ]) ⋅⇩D (I⇩D.map μ ⋆⇩D src⇩D μ)
= (trg⇩D μ ⋆⇩D FG.map μ) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G (D.dom μ))] ⋅⇩D η (D.dom μ) ⋅⇩D 𝗋⇩D[D.dom μ]"
proof -
fix μ
assume μ: "D.arr μ"
have "(𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D η (D.cod μ) ⋅⇩D 𝗋⇩D[D.cod μ]) ⋅⇩D
(I⇩D.map μ ⋆⇩D src⇩D μ)
= 𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D η μ ⋅⇩D 𝗋⇩D[D.dom μ]"
proof -
have "(𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D η (D.cod μ) ⋅⇩D 𝗋⇩D[D.cod μ]) ⋅⇩D
(I⇩D.map μ ⋆⇩D src⇩D μ)
= 𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D (η (D.cod μ) ⋅⇩D μ) ⋅⇩D 𝗋⇩D[D.dom μ]"
using μ D.runit_naturality D.comp_assoc by simp
also have "... = 𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D η μ ⋅⇩D 𝗋⇩D[D.dom μ]"
using μ η_naturality(1) by simp
finally show ?thesis by blast
qed
also have "... = (trg⇩D μ ⋆⇩D FG.map μ) ⋅⇩D 𝗅⇩D⇧-⇧1[FG.map (D.dom μ)] ⋅⇩D η (D.dom μ) ⋅⇩D
𝗋⇩D[D.dom μ]"
proof -
have "𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D η μ ⋅⇩D 𝗋⇩D[D.dom μ] =
𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D (FG.map μ ⋅⇩D η (D.dom μ)) ⋅⇩D 𝗋⇩D[D.dom μ]"
using μ η_naturality(2) D.comp_assoc by simp
also have "... = (𝗅⇩D⇧-⇧1[FG.map (D.cod μ)] ⋅⇩D FG.map μ) ⋅⇩D η (D.dom μ) ⋅⇩D
𝗋⇩D[D.dom μ]"
using D.comp_assoc by simp
also have "... = ((trg⇩D μ ⋆⇩D FG.map μ) ⋅⇩D 𝗅⇩D⇧-⇧1[FG.map (D.dom μ)]) ⋅⇩D
η (D.dom μ) ⋅⇩D 𝗋⇩D[D.dom μ]"
using μ D.lunit'_naturality [of "FG.map μ"] G⇩0_props by simp
also have "... = (trg⇩D μ ⋆⇩D FG.map μ) ⋅⇩D 𝗅⇩D⇧-⇧1[FG.map (D.dom μ)] ⋅⇩D
η (D.dom μ) ⋅⇩D 𝗋⇩D[D.dom μ]"
using D.comp_assoc by simp
finally show ?thesis by blast
qed
finally show "(𝗅⇩D⇧-⇧1[F (G (D.cod μ))] ⋅⇩D η (D.cod μ) ⋅⇩D 𝗋⇩D[D.cod μ]) ⋅⇩D
(I⇩D.map μ ⋆⇩D src⇩D μ)
= (trg⇩D μ ⋆⇩D FG.map μ) ⋅⇩D 𝗅⇩D⇧-⇧1[F (G (D.dom μ))] ⋅⇩D η (D.dom μ) ⋅⇩D
𝗋⇩D[D.dom μ]"
by simp
qed
show "⋀f g. ⟦D.ide f; D.ide g; src⇩D g = trg⇩D f⟧
⟹ (trg⇩D g ⋆⇩D FG.cmp (g, f)) ⋅⇩D
𝖺⇩D[trg⇩D g, FG.map g, FG.map f] ⋅⇩D
(𝗅⇩D⇧-⇧1[F (G g)] ⋅⇩D η g ⋅⇩D 𝗋⇩D[g] ⋆⇩D FG.map f) ⋅⇩D
D.inv 𝖺⇩D[I⇩D.map g, src⇩D g, FG.map f] ⋅⇩D
(I⇩D.map g ⋆⇩D 𝗅⇩D⇧-⇧1[F (G f)] ⋅⇩D η f ⋅⇩D 𝗋⇩D[f]) ⋅⇩D
𝖺⇩D[I⇩D.map g, I⇩D.map f, src⇩D f]
= (𝗅⇩D⇧-⇧1[F (G (g ⋆⇩D f))] ⋅⇩D η (g ⋆⇩D f) ⋅⇩D
𝗋⇩D[g ⋆⇩D f]) ⋅⇩D
(I⇩D.cmp (g, f) ⋆⇩D src⇩D f)"
using η_respects_comp by simp
show "⋀a. D.obj a ⟹ (a ⋆⇩D FG.unit a) ⋅⇩D 𝗋⇩D⇧-⇧1[a] ⋅⇩D 𝗅⇩D[a] =
(𝗅⇩D⇧-⇧1[F (G a)] ⋅⇩D η a ⋅⇩D 𝗋⇩D[a]) ⋅⇩D (I⇩D.unit a ⋆⇩D a)"
using η_respects_unit by auto
show "⋀a. D.obj a ⟹ D.equivalence_map a"
using D.obj_is_equivalence_map by simp
qed
abbreviation unit⇩0'
where "unit⇩0' ≡ λa. a"
abbreviation unit⇩1'
where "unit⇩1' ≡ λf. 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]"
interpretation η: pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
I⇩C.map I⇩C.cmp GF.map GF.cmp unit⇩0' unit⇩1'
proof
show "⋀a. C.obj a ⟹ C.ide a"
by auto
show "⋀f. C.ide f ⟹ C.iso (𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)])"
using C.iso_runit C.iso_lunit'
by (intro C.isos_compose) auto
show "⋀a. C.obj a ⟹ «a : src⇩C (I⇩C.map a) →⇩C src⇩C (GF.map a)»"
by (simp_all add: C.obj_simps(1-3) G⇩0_props(2))
show "⋀f. C.ide f ⟹ «𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)] :
GF.map f ⋆⇩C src⇩C f ⇒⇩C trg⇩C f ⋆⇩C I⇩C.map f»"
using G⇩0_props by auto
show "⋀μ. C.arr μ ⟹
(𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε (C.cod μ) ⋅⇩C 𝗋⇩C[G (F (C.cod μ))]) ⋅⇩C
(GF.map μ ⋆⇩C src⇩C μ)
= (trg⇩C μ ⋆⇩C I⇩C.map μ) ⋅⇩C 𝗅⇩C⇧-⇧1[C.dom μ] ⋅⇩C ε (C.dom μ) ⋅⇩C
𝗋⇩C[G (F (C.dom μ))]"
proof -
fix μ
assume μ: "C.arr μ"
have "(𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε (C.cod μ) ⋅⇩C 𝗋⇩C[GF.map (C.cod μ)]) ⋅⇩C
(GF.map μ ⋆⇩C src⇩C μ)
= 𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε μ ⋅⇩C 𝗋⇩C[GF.map (C.dom μ)]"
proof -
have "(𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε (C.cod μ) ⋅⇩C 𝗋⇩C[GF.map (C.cod μ)]) ⋅⇩C
(GF.map μ ⋆⇩C src⇩C μ)
= 𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε (C.cod μ) ⋅⇩C 𝗋⇩C[GF.map (C.cod μ)] ⋅⇩C
(GF.map μ ⋆⇩C src⇩C μ)"
using C.comp_assoc by simp
also have "... = 𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C (ε (C.cod μ) ⋅⇩C GF.map μ) ⋅⇩C 𝗋⇩C[GF.map (C.dom μ)]"
using μ C.runit_naturality [of "GF.map μ"] G⇩0_props C.comp_assoc by simp
also have "... = 𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε μ ⋅⇩C 𝗋⇩C[GF.map (C.dom μ)]"
using μ ε_naturality(1) [of μ] by simp
finally show ?thesis by blast
qed
also have "... = (trg⇩C μ ⋆⇩C I⇩C.map μ) ⋅⇩C 𝗅⇩C⇧-⇧1[C.dom μ] ⋅⇩C ε (C.dom μ) ⋅⇩C
𝗋⇩C[GF.map (C.dom μ)]"
proof -
have "... = 𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C (μ ⋅⇩C ε (C.dom μ)) ⋅⇩C 𝗋⇩C[GF.map (C.dom μ)]"
using μ ε_naturality(2) [of μ] by simp
also have "... = (𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C μ) ⋅⇩C ε (C.dom μ) ⋅⇩C 𝗋⇩C[GF.map (C.dom μ)]"
using C.comp_assoc by simp
also have "... = ((trg⇩C μ ⋆⇩C I⇩C.map μ) ⋅⇩C 𝗅⇩C⇧-⇧1[C.dom μ]) ⋅⇩C ε (C.dom μ) ⋅⇩C
𝗋⇩C[GF.map (C.dom μ)]"
using μ C.lunit'_naturality [of μ] by simp
also have "... = (trg⇩C μ ⋆⇩C I⇩C.map μ) ⋅⇩C 𝗅⇩C⇧-⇧1[C.dom μ] ⋅⇩C ε (C.dom μ) ⋅⇩C
𝗋⇩C[GF.map (C.dom μ)]"
using C.comp_assoc by simp
finally show ?thesis by blast
qed
finally show "(𝗅⇩C⇧-⇧1[C.cod μ] ⋅⇩C ε (C.cod μ) ⋅⇩C 𝗋⇩C[G (F (C.cod μ))]) ⋅⇩C
(GF.map μ ⋆⇩C src⇩C μ)
= (trg⇩C μ ⋆⇩C I⇩C.map μ) ⋅⇩C 𝗅⇩C⇧-⇧1[C.dom μ] ⋅⇩C ε (C.dom μ) ⋅⇩C
𝗋⇩C[G (F (C.dom μ))]"
by simp
qed
show "⋀f g. ⟦C.ide f; C.ide g; src⇩C g = trg⇩C f⟧ ⟹
(trg⇩C g ⋆⇩C I⇩C.cmp (g, f)) ⋅⇩C
𝖺⇩C[trg⇩C g, I⇩C.map g, I⇩C.map f] ⋅⇩C
(𝗅⇩C⇧-⇧1[g] ⋅⇩C ε g ⋅⇩C 𝗋⇩C[G (F g)] ⋆⇩C I⇩C.map f) ⋅⇩C
C.inv 𝖺⇩C[GF.map g, src⇩C g, I⇩C.map f] ⋅⇩C
(GF.map g ⋆⇩C 𝗅⇩C⇧-⇧1[f] ⋅⇩C ε f ⋅⇩C 𝗋⇩C[G (F f)]) ⋅⇩C
𝖺⇩C[GF.map g, GF.map f, src⇩C f]
= (𝗅⇩C⇧-⇧1[g ⋆⇩C f] ⋅⇩C ε (g ⋆⇩C f) ⋅⇩C 𝗋⇩C[G (F (g ⋆⇩C f))]) ⋅⇩C
(GF.cmp (g, f) ⋆⇩C src⇩C f)"
using ε_respects_comp by simp
show "⋀a. C.obj a ⟹ (a ⋆⇩C I⇩C.unit a) ⋅⇩C 𝗋⇩C⇧-⇧1[a] ⋅⇩C 𝗅⇩C[a] =
(𝗅⇩C⇧-⇧1[a] ⋅⇩C ε a ⋅⇩C 𝗋⇩C[G (F a)]) ⋅⇩C (GF.unit a ⋆⇩C a)"
using ε_respects_unit by auto
show "⋀a. C.obj a ⟹ C.equivalence_map a"
using C.obj_is_equivalence_map by simp
qed
interpretation EQ: equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F Φ G Φ⇩G.map unit⇩0' unit⇩1' counit⇩0' counit⇩1'
..
lemma extends_to_equivalence_of_bicategories:
shows "equivalence_of_bicategories V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F Φ G Φ⇩G.map unit⇩0' unit⇩1' counit⇩0' counit⇩1'"
..
end
subsection "Equivalence Pseudofunctors Extend to Equivalences of Bicategories"
text ‹
Now we put the pieces together and prove that an arbitrary equivalence pseudofunctor extends
to an equivalence of bicategories.
›
context equivalence_pseudofunctor
begin
text ‹
Define a set of objects ‹U› of ‹C› by choosing a representative of each equivalence
class of objects having the same image under the object map of the given equivalence
pseudofunctor. Then ‹U› is obviously dense, because every object of ‹C› belongs to
such an equivalence class.
›
definition U
where "U = {a. C.obj a ∧ a = (SOME a'. C.obj a' ∧ map⇩0 a' = map⇩0 a)}"
lemma U_dense:
assumes "C.obj a"
shows "∃a' ∈ U. C.equivalent_objects a a'"
proof -
let ?a' = "SOME a'. C.obj a' ∧ map⇩0 a' = map⇩0 a"
have "∃a'. C.obj a' ∧ map⇩0 a' = map⇩0 a"
using assms by auto
hence 1: "?a' ∈ U ∧ map⇩0 ?a' = map⇩0 a"
using assms U_def someI_ex [of "λa'. C.obj a' ∧ map⇩0 a' = map⇩0 a"] by simp
hence "C.equivalent_objects ?a' a"
using D.equivalent_objects_reflexive reflects_equivalent_objects [of ?a' a]
by (simp add: U_def assms)
thus ?thesis
using 1 C.equivalent_objects_symmetric by auto
qed
text ‹
Take ‹V› to be the collection of images of all objects of ‹C› under the given equivalence
pseudofunctor. Since equivalence pseudofunctors are biessentially surjective on objects,
‹V› is dense. Moreover, by construction, the object map of the given equivalence
pseudofunctor is a bijection from ‹U› to ‹V›.
›
definition V
where "V = map⇩0 ` Collect C.obj"
lemma V_dense:
assumes "D.obj b"
shows "∃b'. b' ∈ map⇩0 ` Collect C.obj ∧ D.equivalent_objects b b'"
using assms biessentially_surjective_on_objects D.equivalent_objects_symmetric
by blast
lemma bij_betw_U_V:
shows "bij_betw map⇩0 U V"
proof -
have "inj_on map⇩0 U"
using U_def by (intro inj_onI) simp
moreover have "map⇩0 ` U = V"
proof
show "map⇩0 ` U ⊆ V"
using U_def V_def by auto
show "V ⊆ map⇩0 ` U"
proof
fix b
assume b: "b ∈ V"
obtain a where a: "C.obj a ∧ map⇩0 a = b"
using b V_def by auto
let ?a' = "SOME a'. C.obj a' ∧ map⇩0 a' = map⇩0 a"
have 1: "C.obj ?a' ∧ map⇩0 ?a' = b"
using a someI_ex [of "λa'. C.obj a' ∧ map⇩0 a' = map⇩0 a"] by auto
moreover have "?a' = (SOME a''. C.obj a'' ∧ map⇩0 a'' = map⇩0 ?a')"
using a 1 by simp
ultimately have "?a' ∈ U"
unfolding U_def by simp
thus "b ∈ map⇩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⇩0 U V] by simp
qed
abbreviation (input) Arr⇩U
where "Arr⇩U ≡ λμ. C.arr μ ∧ src⇩C μ ∈ U ∧ trg⇩C μ ∈ U"
interpretation C⇩U: subbicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C Arr⇩U
using C.𝔩_ide_simp C.𝔯_ide_simp
apply unfold_locales
apply auto
apply (metis C.comp_ide_self C.ide_src C.src_cod C.src_dom)
by (metis C.trg.as_nat_trans.naturality2 C.trg.as_nat_trans.naturality C.trg_cod)
interpretation C⇩U: dense_subbicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C U
proof
show "⋀a. C.obj a ⟹ ∃a'. a' ∈ U ∧ C.equivalent_objects a' a"
using U_dense C.equivalent_objects_symmetric by blast
qed
abbreviation (input) Arr⇩V
where "Arr⇩V ≡ λμ. D.arr μ ∧ src⇩D μ ∈ V ∧ trg⇩D μ ∈ V"
interpretation D⇩V: subbicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D Arr⇩V
using D.𝔩_ide_simp D.𝔯_ide_simp
apply unfold_locales
apply auto
apply (metis D.comp_ide_self D.ide_src D.src_cod D.src_dom)
by (metis D.trg.as_nat_trans.naturality2 D.trg.as_nat_trans.naturality D.trg_cod)
interpretation D⇩V: dense_subbicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V
using V_dense D.equivalent_objects_def D.equivalent_objects_symmetric V_def
by unfold_locales metis
interpretation F⇩U: restricted_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ
‹λμ. C.arr μ ∧ src⇩C μ ∈ U ∧ trg⇩C μ ∈ U›
..
interpretation F⇩U⇩V: corestricted_pseudofunctor
C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F⇩U.map F⇩U.cmp ‹λa. a ∈ V›
proof
show "⋀μ. C⇩U.arr μ ⟹ D⇩V.arr (F⇩U.map μ)"
proof -
fix μ
assume μ: "C⇩U.arr μ"
have "src⇩C μ ∈ U ∧ trg⇩C μ ∈ U"
using μ C⇩U.arr_char⇩S⇩b⇩C by simp
moreover have "⋀a. a ∈ U ⟹ F⇩U.map⇩0 a = map⇩0 a"
proof -
fix a :: 'a
assume a: "a ∈ U"
hence 1: "C.obj a"
using U_def by blast
have "src⇩C a = a"
using a U_def by blast
thus "F⇩U.map⇩0 a = map⇩0 a"
using a 1 C.obj_simps(1,3) C⇩U.arr_char⇩S⇩b⇩C F⇩U.map⇩0_def map⇩0_def by presburger
qed
ultimately have "F⇩U.map⇩0 (src⇩C μ) ∈ V ∧ F⇩U.map⇩0 (trg⇩C μ) ∈ V"
using μ bij_betw_U_V bij_betw_def
by (metis (no_types, lifting) image_eqI)
hence "src⇩D (F⇩U.map μ) ∈ V ∧ trg⇩D (F⇩U.map μ) ∈ V"
using μ F⇩U.map⇩0_def C⇩U.src_def C⇩U.trg_def F⇩U.preserves_src F⇩U.preserves_trg
by auto
thus "D⇩V.arr (F⇩U.map μ)"
using μ D⇩V.arr_char⇩S⇩b⇩C [of "F⇩U.map μ"] by blast
qed
qed
interpretation F⇩U⇩V: equivalence_pseudofunctor
C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg
F⇩U⇩V.map F⇩U⇩V.cmp
proof
show "⋀f f'. ⟦C⇩U.par f f'; F⇩U.map f = F⇩U.map f'⟧ ⟹ f = f'"
using C⇩U.arr_char⇩S⇩b⇩C C⇩U.dom_char⇩S⇩b⇩C C⇩U.cod_char⇩S⇩b⇩C
by (metis (no_types, lifting) is_faithful)
show "⋀a b g. ⟦C⇩U.obj a; C⇩U.obj b; D⇩V.in_hhom g (F⇩U⇩V.map⇩0 a) (F⇩U⇩V.map⇩0 b);
D⇩V.ide g⟧
⟹ ∃f. C⇩U.in_hhom f a b ∧ C⇩U.ide f ∧ D⇩V.isomorphic (F⇩U.map f) g"
proof -
fix a b g
assume a: "C⇩U.obj a" and b: "C⇩U.obj b"
assume g: "D⇩V.in_hhom g (F⇩U⇩V.map⇩0 a) (F⇩U⇩V.map⇩0 b)" and ide_g: "D⇩V.ide g"
have 1: "∃f. «f : a →⇩C b» ∧ C.ide f ∧ D.isomorphic (F f) g"
proof -
have "C.obj a ∧ C.obj b"
using a b C⇩U.obj_char by simp
moreover have "«g : map⇩0 a →⇩D map⇩0 b»"
using a b g D⇩V.in_hhom_def D⇩V.arr_char⇩S⇩b⇩C D⇩V.src_def D⇩V.trg_def
by (intro D.in_hhomI) auto
moreover have "D.ide g"
using ide_g D⇩V.ide_char⇩S⇩b⇩C D⇩V.arr_char⇩S⇩b⇩C by simp
ultimately show ?thesis
using locally_essentially_surjective by simp
qed
obtain f where f: "«f : a →⇩C b» ∧ C.ide f ∧ D.isomorphic (F f) g"
using 1 by blast
have 2: "C⇩U.in_hhom f a b"
using f a b C⇩U.arr_char⇩S⇩b⇩C C⇩U.obj_char C⇩U.src_def C⇩U.trg_def by fastforce
moreover have "C⇩U.ide f"
using f 2 C⇩U.ide_char⇩S⇩b⇩C C⇩U.arr_char⇩S⇩b⇩C by auto
moreover have "D⇩V.isomorphic (F⇩U⇩V.map f) g"
proof -
obtain φ where φ: "«φ : F f ⇒⇩D g» ∧ D.iso φ"
using f D.isomorphic_def by auto
have 3: "D⇩V.in_hom φ (F⇩U⇩V.map f) g"
proof -
have "D⇩V.arr φ"
using f g φ 2 D⇩V.arr_char⇩S⇩b⇩C
by (metis (no_types, lifting) D.arrI D.vconn_implies_hpar(1-4) D⇩V.ideD(1) ide_g)
thus ?thesis
using φ 2 D⇩V.dom_char⇩S⇩b⇩C D⇩V.cod_char⇩S⇩b⇩C
by (intro D⇩V.in_homI) auto
qed
moreover have "D⇩V.iso φ"
using φ D⇩V.iso_char⇩S⇩b⇩C D⇩V.arr_char⇩S⇩b⇩C 3 by fastforce
ultimately show ?thesis
using D⇩V.isomorphic_def by auto
qed
ultimately show "∃f. C⇩U.in_hhom f a b ∧ C⇩U.ide f ∧ D⇩V.isomorphic (F⇩U⇩V.map f) g"
by auto
qed
show "⋀f f' ν. ⟦C⇩U.ide f; C⇩U.ide f'; C⇩U.src f = C⇩U.src f'; C⇩U.trg f = C⇩U.trg f';
D⇩V.in_hom ν (F⇩U.map f) (F⇩U.map f')⟧
⟹ ∃μ. C⇩U.in_hom μ f f' ∧ F⇩U.map μ = ν"
proof -
fix f f' ν
assume f: "C⇩U.ide f" and f': "C⇩U.ide f'"
and eq_src: "C⇩U.src f = C⇩U.src f'" and eq_trg: "C⇩U.trg f = C⇩U.trg f'"
and ν: "D⇩V.in_hom ν (F⇩U.map f) (F⇩U.map f')"
have 1: "C.ide f ∧ C.ide f'"
using f f' C⇩U.ide_char⇩S⇩b⇩C C⇩U.arr_char⇩S⇩b⇩C by simp
have 2: "∃μ. «μ : f ⇒⇩C f'» ∧ F μ = ν"
proof -
have "src⇩C f = src⇩C f' ∧ trg⇩C f = trg⇩C f'"
using f f' 1 eq_src eq_trg C⇩U.src_def C⇩U.trg_def by simp
moreover have "«ν : F f ⇒⇩D F f'»"
using ν D⇩V.in_hom_char⇩S⇩b⇩C D⇩V.arr_char⇩S⇩b⇩C by auto
ultimately show ?thesis
using 1 locally_full by simp
qed
obtain μ where μ: "«μ : f ⇒⇩C f'» ∧ F μ = ν"
using 2 by auto
have 3: "C⇩U.in_hom μ f f'"
using μ f f' C⇩U.in_hom_char⇩S⇩b⇩C C⇩U.ide_char⇩S⇩b⇩C C⇩U.arr_char⇩S⇩b⇩C by auto
moreover have "F⇩U.map μ = ν"
using μ 3 by auto
ultimately show "∃μ. C⇩U.in_hom μ f f' ∧ F⇩U.map μ = ν"
by auto
qed
show "⋀b. D⇩V.obj b ⟹ ∃a. C⇩U.obj a ∧ D⇩V.equivalent_objects (F⇩U⇩V.map⇩0 a) b"
proof -
fix b
assume b: "D⇩V.obj b"
obtain a where a: "C.obj a ∧ map⇩0 a = b"
using b D⇩V.obj_char D⇩V.arr_char⇩S⇩b⇩C V_def by auto
have 1: "D.obj b ∧ b ∈ V"
using a b D⇩V.obj_char V_def by auto
obtain a' where a': "a' ∈ U ∧ C.equivalent_objects a' a"
using a U_dense C.equivalent_objects_symmetric by blast
have obj_a': "C⇩U.obj a'"
using a' U_def C⇩U.obj_char C⇩U.arr_char⇩S⇩b⇩C by fastforce
moreover have "D⇩V.equivalent_objects (F⇩U⇩V.map⇩0 a') b"
proof -
have "D.equivalent_objects (map⇩0 a') (map⇩0 a)"
using a' preserves_equivalent_objects by simp
hence 2: "D.equivalent_objects (map⇩0 a') b"
using a D.equivalent_objects_transitive by simp
obtain e where e: "«e : map⇩0 a' →⇩D b» ∧ D.equivalence_map e"
using 2 D.equivalent_objects_def by auto
have 3: "D.obj (map⇩0 a') ∧ map⇩0 a' ∈ V"
using a' e U_def V_def by simp
moreover have e_in_hhom: "D⇩V.in_hhom e (F⇩U⇩V.map⇩0 a') b"
proof
show 4: "D⇩V.arr e"
using 1 3 a e b D⇩V.arr_char⇩S⇩b⇩C D⇩V.obj_char V_def
by (metis (no_types, lifting) D.in_hhomE)
show "D⇩V.src e = F⇩U⇩V.map⇩0 a'"
using e 4 D⇩V.src_def F⇩U⇩V.map⇩0_def obj_a' map⇩0_def by auto
show "D⇩V.trg e = b"
using e 4 D⇩V.trg_def by auto
qed
moreover have "D⇩V.equivalence_map e"
proof -
obtain d η ε where d: "equivalence_in_bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D e d η ε"
using e D.equivalence_map_def by auto
interpret e: equivalence_in_bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D e d η ε
using d by simp
have "equivalence_in_bicategory D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg e d η ε"
proof
show ide_e: "D⇩V.ide e"
using e e_in_hhom D⇩V.ide_char⇩S⇩b⇩C by auto
show ide_d: "D⇩V.ide d"
using d ide_e e.antipar D⇩V.ide_char⇩S⇩b⇩C D⇩V.arr_char⇩S⇩b⇩C by simp
show 4: "D⇩V.in_hom η (D⇩V.src e) (D⇩V.hcomp d e)"
proof -
have "D⇩V.hseq d e"
using ide_d ide_e e.antipar D⇩V.src_def D⇩V.trg_def by simp
thus ?thesis
using ide_d ide_e e.unit_in_hom D⇩V.arr_char⇩S⇩b⇩C D⇩V.ide_char⇩S⇩b⇩C
D⇩V.dom_char⇩S⇩b⇩C D⇩V.cod_char⇩S⇩b⇩C D⇩V.src_def D⇩V.trg_def
by (intro D⇩V.in_homI) auto
qed
show 5: "D⇩V.in_hom ε (D⇩V.hcomp e d) (D⇩V.src d)"
proof -
have "D⇩V.hseq e d"
using ide_d ide_e e.antipar D⇩V.src_def D⇩V.trg_def by simp
thus ?thesis
using ide_d ide_e e.antipar e.counit_in_hom D⇩V.arr_char⇩S⇩b⇩C D⇩V.ide_char⇩S⇩b⇩C
D⇩V.dom_char⇩S⇩b⇩C D⇩V.cod_char⇩S⇩b⇩C D⇩V.src_def D⇩V.trg_def
by (intro D⇩V.in_homI) auto
qed
show "D⇩V.iso η"
using 4 D⇩V.iso_char⇩S⇩b⇩C D⇩V.arr_char⇩S⇩b⇩C by fastforce
show "D⇩V.iso ε"
using 5 D⇩V.iso_char⇩S⇩b⇩C D⇩V.arr_char⇩S⇩b⇩C by fastforce
qed
thus ?thesis
using D⇩V.equivalence_map_def by auto
qed
ultimately show ?thesis
using D⇩V.equivalent_objects_def by auto
qed
ultimately show "∃a. C⇩U.obj a ∧ D⇩V.equivalent_objects (F⇩U⇩V.map⇩0 a) b" by auto
qed
qed
interpretation F⇩U⇩V: equivalence_pseudofunctor_bij_on_obj
C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg
F⇩U⇩V.map F⇩U⇩V.cmp
proof
have "Collect C⇩U.obj = U"
using C⇩U.obj_char C⇩U.arr_char⇩S⇩b⇩C U_def by fastforce
moreover have "Collect D⇩V.obj = V"
using D⇩V.obj_char D⇩V.arr_char⇩S⇩b⇩C V_def D.obj_def' map⇩0_simps(1) by auto
moreover have "⋀a. a ∈ Collect C⇩U.obj ⟹ F⇩U⇩V.map⇩0 a = map⇩0 a"
using F⇩U⇩V.map⇩0_def C⇩U.obj_char C⇩U.arr_char⇩S⇩b⇩C D⇩V.src_def
F⇩U.map⇩0_simp F⇩U⇩V.map⇩0_simp by auto
ultimately show "bij_betw F⇩U⇩V.map⇩0 (Collect C⇩U.obj) (Collect D⇩V.obj)"
using bij_betw_U_V
by (simp add: bij_betw_U_V bij_betw_cong)
qed
interpretation EQ⇩U⇩V: equivalence_of_bicategories
D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg
C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
F⇩U⇩V.map F⇩U⇩V.cmp F⇩U⇩V.G F⇩U⇩V.Φ⇩G
F⇩U⇩V.unit⇩0' F⇩U⇩V.unit⇩1' F⇩U⇩V.counit⇩0' F⇩U⇩V.counit⇩1'
using F⇩U⇩V.extends_to_equivalence_of_bicategories by blast
text ‹
Now compose ‹EQ⇩U⇩V› with the equivalence from ‹D⇩V› to ‹D› and the converse of the equivalence
from ‹C⇩U› to ‹C›. The result is an equivalence of bicategories from ‹C› to ‹D›.
›
interpretation EQ⇩C: equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
C⇩U.E C⇩U.Φ⇩E C⇩U.P C⇩U.Φ⇩P
C⇩U.unit⇩0 C⇩U.unit⇩1 C⇩U.counit⇩0 C⇩U.counit⇩1
using C⇩U.induces_equivalence by simp
interpretation EQ⇩C': converse_equivalence_of_bicategories
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
C⇩U.E C⇩U.Φ⇩E C⇩U.P C⇩U.Φ⇩P
C⇩U.unit⇩0 C⇩U.unit⇩1 C⇩U.counit⇩0 C⇩U.counit⇩1
..
interpretation EQ⇩D: equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg
D⇩V.E D⇩V.Φ⇩E D⇩V.P D⇩V.Φ⇩P
D⇩V.unit⇩0 D⇩V.unit⇩1 D⇩V.counit⇩0 D⇩V.counit⇩1
using D⇩V.induces_equivalence by simp
interpretation EQ⇩U⇩VoEQ⇩C': composite_equivalence_of_bicategories
D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg
C⇩U.comp C⇩U.hcomp C⇩U.𝖺 𝗂⇩C C⇩U.src C⇩U.trg
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F⇩U⇩V.map F⇩U⇩V.cmp F⇩U⇩V.G F⇩U⇩V.Φ⇩G
C⇩U.P C⇩U.Φ⇩P C⇩U.E C⇩U.Φ⇩E
F⇩U⇩V.unit⇩0' F⇩U⇩V.unit⇩1' F⇩U⇩V.counit⇩0' F⇩U⇩V.counit⇩1'
EQ⇩C'.unit⇩0 EQ⇩C'.unit⇩1 EQ⇩C'.counit⇩0 EQ⇩C'.counit⇩1
..
interpretation EQ⇩DoEQ⇩U⇩VoEQ⇩C': composite_equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
D⇩V.comp D⇩V.hcomp D⇩V.𝖺 𝗂⇩D D⇩V.src D⇩V.trg
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
D⇩V.E D⇩V.Φ⇩E D⇩V.P D⇩V.Φ⇩P
EQ⇩U⇩VoEQ⇩C'.left_map EQ⇩U⇩VoEQ⇩C'.left_cmp
EQ⇩U⇩VoEQ⇩C'.right_map EQ⇩U⇩VoEQ⇩C'.right_cmp
D⇩V.unit⇩0 D⇩V.unit⇩1 D⇩V.counit⇩0 D⇩V.counit⇩1
EQ⇩U⇩VoEQ⇩C'.unit⇩0 EQ⇩U⇩VoEQ⇩C'.unit⇩1
EQ⇩U⇩VoEQ⇩C'.counit⇩0 EQ⇩U⇩VoEQ⇩C'.counit⇩1
..
lemma induces_equivalence_of_bicategories:
shows "equivalence_of_bicategories V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩1
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩1"
..
lemma left_map_simp:
assumes "C.arr μ"
shows "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map μ = D⇩V.E (F (C⇩U.P μ))"
using assms by simp
lemma right_map_simp:
assumes "D.arr ν"
shows "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map ν = C⇩U.E (F⇩U⇩V.G (D⇩V.P ν))"
using assms by simp
lemma unit⇩0_simp:
assumes "C.obj a"
shows "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩0 a =
C⇩U.E (F⇩U⇩V.G (D⇩V.η⇩0 (D⇩V.src (F (C⇩U.P a))))) ⋆⇩C C⇩U.E (C⇩U.P⇩0 (src⇩C a))
⋆⇩C EQ⇩C'.unit⇩0 a"
using assms EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩0_simp EQ⇩U⇩VoEQ⇩C'.unit⇩0_simp
EQ⇩U⇩VoEQ⇩C'.FH.map⇩0_def C⇩U.src_def
by auto
text ‹
We've now got an equivalence of bicategories between ‹C› and ‹D›, but it involves
‹EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map› and not the originally given equivalence pseudofunctor ‹F›.
However, we can patch things up by showing that ‹EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map› is pseudonaturally
equivalent to ‹F›. From this, we may conclude, using the fact that equivalences of
bicategories respect pseudonatural equivalence, that there is an equivalence of bicategories
between ‹C› and ‹D› that involves ‹F› and ‹EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map›, rather than
‹EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map› and ‹EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map›.
›
abbreviation τ⇩0
where "τ⇩0 a ≡ F (C⇩U.ε⇩0 a)"
abbreviation τ⇩1
where "τ⇩1 f ≡ D.inv (Φ (C⇩U.ε⇩0 (trg⇩C f), C⇩U.P f)) ⋅⇩D F (C⇩U.ε⇩1 f) ⋅⇩D Φ (f, C⇩U.ε⇩0 (src⇩C f))"
lemma τ⇩0_in_hom [intro]:
assumes "C.obj a"
shows "«τ⇩0 a : map⇩0 (C⇩U.P⇩0 a) →⇩D map⇩0 a»"
and "«τ⇩0 a : τ⇩0 a ⇒⇩D τ⇩0 a»"
proof -
show "«τ⇩0 a : map⇩0 (C⇩U.P⇩0 a) →⇩D map⇩0 a»"
proof
show "D.arr (τ⇩0 a)"
using assms by simp
show "src⇩D (τ⇩0 a) = map⇩0 (C⇩U.P⇩0 a)"
using assms map⇩0_def C⇩U.counit.ide_map⇩0_obj C⇩U.equivalence_data_simps⇩B(7) C.ideD(1)
preserves_src
by presburger
show "trg⇩D (τ⇩0 a) = map⇩0 a"
using assms by auto
qed
show "«τ⇩0 a : τ⇩0 a ⇒⇩D τ⇩0 a»"
using assms by auto
qed
lemma τ⇩0_simps [simp]:
assumes "C.obj a"
shows "D.ide (τ⇩0 a)"
and "src⇩D (τ⇩0 a) = map⇩0 (C⇩U.P⇩0 a)" and "trg⇩D (τ⇩0 a) = map⇩0 a"
using assms τ⇩0_in_hom(1) [of a] by auto
lemma τ⇩1_in_hom [intro]:
assumes "C.ide f"
shows "«τ⇩1 f : map⇩0 (C⇩U.P⇩0 (src⇩C f)) →⇩D map⇩0 (trg⇩C f)»"
and "«τ⇩1 f : F f ⋆⇩D τ⇩0 (src⇩C f) ⇒⇩D τ⇩0 (trg⇩C f) ⋆⇩D F (C⇩U.P f)»"
proof -
show 1: "«τ⇩1 f : F f ⋆⇩D τ⇩0 (src⇩C f) ⇒⇩D τ⇩0 (trg⇩C f) ⋆⇩D F (C⇩U.P f)»"
proof (intro D.comp_in_homI)
show "«Φ (f, C⇩U.d (src⇩C f)) : F f ⋆⇩D τ⇩0 (src⇩C f) ⇒⇩D F (f ⋆⇩C C⇩U.ε⇩0 (src⇩C f))»"
using assms by auto
show "«F (C⇩U.ε⇩1 f) : F (f ⋆⇩C C⇩U.ε⇩0 (src⇩C f)) ⇒⇩D F (C⇩U.ε⇩0 (trg⇩C f) ⋆⇩C C⇩U.P f)»"
using assms C⇩U.counit⇩1_in_hom [of f] C⇩U.P_def by auto
show "«D.inv (Φ (C⇩U.ε⇩0 (trg⇩C f), C⇩U.P f)) :
F (C⇩U.ε⇩0 (trg⇩C f) ⋆⇩C C⇩U.P f) ⇒⇩D τ⇩0 (trg⇩C f) ⋆⇩D F (C⇩U.P f)»"
using assms C.VV.arr_char⇩S⇩b⇩C C.VV.ide_char⇩S⇩b⇩C Φ.components_are_iso
by (metis (no_types, lifting) C⇩U.P_def C⇩U.counit.ide_map⇩0_obj C⇩U.counit⇩1_simps(1,5)
C⇩U.equivalence_data_simps⇩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 "«τ⇩1 f : map⇩0 (C⇩U.P⇩0 (src⇩C f)) →⇩D map⇩0 (trg⇩C f)»"
using assms 1 map⇩0_def [of "C⇩U.P⇩0 (src⇩C f)"] C⇩U.emb.map⇩0_simp C⇩U.P⇩0_props
apply (intro D.in_hhomI)
apply auto
using D.src_dom [of "τ⇩1 f"]
apply (elim D.in_homE)
apply auto
using D.trg_dom [of "τ⇩1 f"]
apply (elim D.in_homE)
by auto
qed
lemma τ⇩1_simps [simp]:
assumes "C.ide f"
shows "D.arr (τ⇩1 f)"
and "src⇩D (τ⇩1 f) = map⇩0 (C⇩U.P⇩0 (src⇩C f))" and "trg⇩D (τ⇩1 f) = map⇩0 (trg⇩C f)"
and "D.dom (τ⇩1 f) = F f ⋆⇩D τ⇩0 (src⇩C f)" and "D.cod (τ⇩1 f) = τ⇩0 (trg⇩C f) ⋆⇩D F (C⇩U.P f)"
using assms τ⇩1_in_hom by blast+
lemma iso_τ⇩1:
assumes "C.ide f"
shows "D.iso (τ⇩1 f)"
proof -
have "C.VV.ide (C⇩U.ε⇩0 (trg⇩C f), C⇩U.P f)"
proof -
have "C.VV.arr (C⇩U.ε⇩0 (trg⇩C f), C⇩U.P f)"
using assms C⇩U.equivalence_data_simps⇩B(7) C.ideD(1) by auto
moreover have "C.VxV.ide (C⇩U.ε⇩0 (trg⇩C f), C⇩U.P f)"
using assms C⇩U.prj.preserves_ide [of f] C⇩U.ide_char⇩S⇩b⇩C by simp
ultimately show ?thesis
using assms C.VV.ide_char⇩S⇩b⇩C [of "(C⇩U.ε⇩0 (trg⇩C f), C⇩U.P f)"] by blast
qed
hence "D.iso (Φ (C⇩U.d (trg⇩C f), C⇩U.P f))"
by simp
thus ?thesis
using assms C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C Φ.components_are_iso
by (intro D.isos_compose) auto
qed
interpretation τ: pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp F Φ
τ⇩0 τ⇩1
proof
show "⋀a. C.obj a ⟹ D.ide (τ⇩0 a)"
by simp
show "⋀a. C.obj a ⟹ D.equivalence_map (τ⇩0 a)"
using C⇩U.counit.components_are_equivalences preserves_equivalence_maps by blast
show "⋀a. C.obj a ⟹ «τ⇩0 a : src⇩D (EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map a) →⇩D src⇩D (F a)»"
proof -
fix a
assume a: "C.obj a"
show "«τ⇩0 a : src⇩D (EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map a) →⇩D src⇩D (F a)»"
proof
show "D.arr (τ⇩0 a)"
using a by simp
show "src⇩D (τ⇩0 a) = src⇩D (EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map a)"
proof -
have "D⇩V.arr (F (C⇩U.P a))"
proof -
have "src⇩D (F (C⇩U.P a)) ∈ V"
proof -
have "src⇩D (F (C⇩U.P a)) = map⇩0 (C⇩U.prj.map⇩0 a)"
using a by auto
moreover have "C.obj (C⇩U.prj.map⇩0 a)"
using a C⇩U.obj_char [of "C⇩U.prj.map⇩0 a"] C⇩U.prj.map⇩0_simps(1) by auto
ultimately show ?thesis
using V_def by simp
qed
moreover have "trg⇩D (F (C⇩U.P a)) ∈ V"
proof -
have "trg⇩D (F (C⇩U.P a)) = map⇩0 (C⇩U.prj.map⇩0 a)"
using a by auto
moreover have "C.obj (C⇩U.prj.map⇩0 a)"
using a C⇩U.obj_char [of "C⇩U.prj.map⇩0 a"] C⇩U.prj.map⇩0_simps(1) by auto
ultimately show ?thesis
using V_def by simp
qed
ultimately show ?thesis
using a D⇩V.arr_char⇩S⇩b⇩C by fastforce
qed
thus ?thesis
using a D⇩V.emb.map⇩0_def D⇩V.emb.map_def D⇩V.src_def C.obj_simps(1-2) C⇩U.P_simps⇩B(2)
by (simp add: C⇩U.P⇩0_props(1))
qed
show "trg⇩D (F (C⇩U.d a)) = src⇩D (F a)"
using a by auto
qed
qed
show "⋀f. C.ide f ⟹
«τ⇩1 f : F f ⋆⇩D τ⇩0 (src⇩C f) ⇒⇩D τ⇩0 (trg⇩C f) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f»"
proof -
fix f
assume f: "C.ide f"
show "«τ⇩1 f : F f ⋆⇩D τ⇩0 (src⇩C f) ⇒⇩D τ⇩0 (trg⇩C f) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f»"
proof -
have "D⇩V.arr (F (C⇩U.P f))"
using f F⇩U⇩V.preserves_arr by auto
hence "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f = F (C⇩U.P f)"
using f EQ⇩U⇩VoEQ⇩C'.FH.map⇩0_def D⇩V.emb.map_def by simp
thus ?thesis
using f by auto
qed
qed
show "⋀f. C.ide f ⟹ D.iso (τ⇩1 f)"
using iso_τ⇩1 by simp
show "⋀μ. C.arr μ ⟹
τ⇩1 (C.cod μ) ⋅⇩D (F μ ⋆⇩D τ⇩0 (src⇩C μ)) =
(τ⇩0 (trg⇩C μ) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map μ) ⋅⇩D τ⇩1 (C.dom μ)"
proof -
fix μ
assume μ: "C.arr μ"
show "τ⇩1 (C.cod μ) ⋅⇩D (F μ ⋆⇩D τ⇩0 (src⇩C μ))
= (τ⇩0 (trg⇩C μ) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map μ) ⋅⇩D τ⇩1 (C.dom μ)"
proof -
have "τ⇩1 (C.cod μ) ⋅⇩D (F μ ⋆⇩D τ⇩0 (src⇩C μ))
= D.inv (Φ (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P (C.cod μ))) ⋅⇩D
F (C⇩U.ε⇩1 (C.cod μ)) ⋅⇩D
Φ (C.cod μ, C⇩U.ε⇩0 (src⇩C μ)) ⋅⇩D
(F μ ⋆⇩D τ⇩0 (src⇩C μ))"
using μ D.comp_assoc by simp
also have "... = D.inv (Φ (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P (C.cod μ))) ⋅⇩D
(F (C⇩U.ε⇩1 (C.cod μ)) ⋅⇩D
F (μ ⋆⇩C C⇩U.ε⇩0 (src⇩C μ))) ⋅⇩D
Φ (C.dom μ, C⇩U.ε⇩0 (src⇩C μ))"
using μ Φ.naturality [of "(μ, C⇩U.ε⇩0 (src⇩C μ))"] D.comp_assoc
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def
by simp
also have "... = (D.inv (Φ (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P (C.cod μ))) ⋅⇩D
F (C⇩U.ε⇩0 (trg⇩C μ) ⋆⇩C C⇩U.EoP.map μ)) ⋅⇩D
F (C⇩U.ε⇩1 (C.dom μ)) ⋅⇩D
Φ (C.dom μ, C⇩U.ε⇩0 (src⇩C μ))"
proof -
have "F (C⇩U.ε⇩1 (C.cod μ)) ⋅⇩D F (μ ⋆⇩C C⇩U.ε⇩0 (src⇩C μ))
= F (C⇩U.ε⇩1 (C.cod μ) ⋅⇩C (μ ⋆⇩C C⇩U.ε⇩0 (src⇩C μ)))"
using μ by simp
also have "... = F ((C⇩U.ε⇩0 (trg⇩C μ) ⋆⇩C C⇩U.EoP.map μ) ⋅⇩C C⇩U.ε⇩1 (C.dom μ))"
using μ C⇩U.counit.naturality [of μ] by simp
also have "... = F (C⇩U.ε⇩0 (trg⇩C μ) ⋆⇩C C⇩U.EoP.map μ) ⋅⇩D
F (C⇩U.ε⇩1 (C.dom μ))"
using μ by simp
finally have "F (C⇩U.ε⇩1 (C.cod μ)) ⋅⇩D F (μ ⋆⇩C C⇩U.ε⇩0 (src⇩C μ))
= F (C⇩U.ε⇩0 (trg⇩C μ) ⋆⇩C C⇩U.EoP.map μ) ⋅⇩D F (C⇩U.ε⇩1 (C.dom μ))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F (C⇩U.ε⇩0 (trg⇩C μ)) ⋆⇩D F (C⇩U.EoP.map μ)) ⋅⇩D
D.inv (Φ (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P (C.dom μ))) ⋅⇩D
F (C⇩U.ε⇩1 (C.dom μ)) ⋅⇩D
Φ (C.dom μ, C⇩U.ε⇩0 (src⇩C μ))"
proof -
have "D.inv (Φ (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P (C.cod μ))) ⋅⇩D
F (C⇩U.ε⇩0 (trg⇩C μ) ⋆⇩C C⇩U.EoP.map μ)
= (F (C⇩U.ε⇩0 (trg⇩C μ)) ⋆⇩D F (C⇩U.EoP.map μ)) ⋅⇩D
D.inv (Φ (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P (C.dom μ)))"
proof -
have "C.VV.arr (C⇩U.ε⇩0 (trg⇩C μ), C⇩U.P μ)"
proof (unfold C.VV.arr_char⇩S⇩b⇩C, intro conjI)
show "C.arr (fst (C⇩U.d (trg⇩C μ), C⇩U.P μ))"
using μ by simp
show "C.arr (snd (C⇩U.d (trg⇩C μ), C⇩U.P μ))"
using μ by simp
show "src⇩C (fst (C⇩U.d (trg⇩C μ), C⇩U.P μ)) = trg⇩C (snd (C⇩U.d (trg⇩C μ), C⇩U.P μ))"
using μ C⇩U.emb.map⇩0_def C⇩U.emb.map_def apply simp
using C⇩U.P⇩0_props(1) C⇩U.P_simps⇩B(3) C⇩U.src_def by fastforce
qed
moreover have "C⇩U.EoP.map μ = C⇩U.P μ"
using μ by (simp add: C⇩U.emb.map_def)
moreover have "C⇩U.emb.map⇩0 (C⇩U.P⇩0 (trg⇩C μ)) = C⇩U.P⇩0 (trg⇩C μ)"
using μ C⇩U.P⇩0_props(1) C⇩U.emb.map⇩0_simp by blast
ultimately show ?thesis
using μ C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C.VV.ide_char⇩S⇩b⇩C FF_def
Φ.inv_naturality [of "(C⇩U.counit⇩0 (trg⇩C μ), C⇩U.P μ)"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (τ⇩0 (trg⇩C μ) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map μ) ⋅⇩D τ⇩1 (C.dom μ)"
using μ F⇩U⇩V.preserves_arr D.comp_assoc D⇩V.emb.map_def C⇩U.emb.map_def by simp
finally show "τ⇩1 (C.cod μ) ⋅⇩D (F μ ⋆⇩D τ⇩0 (src⇩C μ))
= (τ⇩0 (trg⇩C μ) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map μ) ⋅⇩D τ⇩1 (C.dom μ)"
by blast
qed
qed
show "⋀a. C.obj a ⟹
(τ⇩0 a ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.unit a) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a] ⋅⇩D 𝗅⇩D[τ⇩0 a] =
τ⇩1 a ⋅⇩D (unit a ⋆⇩D τ⇩0 a)"
proof -
fix a
assume a: "C.obj a"
have 1: "C⇩U.obj (C⇩U.prj.map⇩0 a)"
using a C⇩U.prj.map⇩0_simps(1) by simp
have 2: "«C⇩U.prj.unit a : C⇩U.prj.map⇩0 a ⇒⇩C C⇩U.P a»"
using a C⇩U.in_hom_char⇩S⇩b⇩C by blast
have 3: "«C⇩U.prj.unit a : C⇩U.prj.map⇩0 a →⇩C C⇩U.prj.map⇩0 a»"
proof (intro C.in_hhomI)
show 4: "C.arr (C⇩U.prj.unit a)"
using 2 by auto
show "src⇩C (C⇩U.prj.unit a) = C⇩U.prj.map⇩0 a"
proof -
have "src⇩C (C⇩U.prj.unit a) = src⇩C (C⇩U.P a)"
using 2 4 C.src_cod [of "C⇩U.prj.unit a"] C.vconn_implies_hpar(1,3) by auto
also have "... = C⇩U.prj.map⇩0 a"
using a C.obj_simps(1) C⇩U.prj.map⇩0_def [of a] C⇩U.src_def [of "C⇩U.P a"]
by simp
finally show ?thesis by blast
qed
show "trg⇩C (C⇩U.prj.unit a) = C⇩U.prj.map⇩0 a"
proof -
have "trg⇩C (C⇩U.prj.unit a) = trg⇩C (C⇩U.P a)"
using 2 4 C.trg_cod [of "C⇩U.prj.unit a"] C.vconn_implies_hpar(2,4) by auto
also have "... = C⇩U.prj.map⇩0 a"
using a C.obj_simps(1,3) C⇩U.prj.map⇩0_def [of a] C⇩U.trg_def [of "C⇩U.P a"]
by simp
finally show ?thesis by blast
qed
qed
have [simp]: "C.dom (C⇩U.prj.unit a) = C⇩U.prj.map⇩0 a"
using 2 by auto
have [simp]: "C.arr (C⇩U.prj.unit a)"
using 2 by auto
have [simp]: "C.cod (C⇩U.prj.unit a) = C⇩U.P a"
using 2 by auto
have [simp]: "src⇩C (C⇩U.prj.unit a) = C⇩U.prj.map⇩0 a"
using 3 by auto
have [simp]: "trg⇩C (C⇩U.prj.unit a) = C⇩U.prj.map⇩0 a"
using 3 by auto
have [simp]: "C.arr (C⇩U.d a)"
using a by simp
have [simp]: "C.ide (C⇩U.d a)"
using a by simp
have [simp]: "C.dom (C⇩U.d a) = C⇩U.d a"
using a by auto
have [simp]: "C.cod (C⇩U.d a) = C⇩U.d a"
using a by auto
have [simp]: "src⇩C (C⇩U.d a) = C⇩U.prj.map⇩0 a"
using a C⇩U.equivalence_data_simps⇩B(7) by auto
have [simp]: "trg⇩C (C⇩U.d a) = a"
using a C⇩U.equivalence_data_simps⇩B(7) by auto
have seq: "D.seq (F (C⇩U.prj.unit a)) (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using a
apply (intro D.seqI)
using C⇩U.prj.map⇩0_simps(1) D⇩V.arr_char⇩S⇩b⇩C F⇩U⇩V.unit_simps(1) apply blast
using C⇩U.prj.unit_simps(1) C⇩U.arr_char⇩S⇩b⇩C apply blast
proof -
have "D.dom (F (C⇩U.prj.unit a)) = F (C.dom (C⇩U.prj.unit a))"
using a C⇩U.prj.unit_simps(1) C⇩U.arr_char⇩S⇩b⇩C by simp
also have "... = F (C⇩U.prj.map⇩0 a)"
using a C⇩U.prj.unit_simps [of a] C⇩U.dom_char⇩S⇩b⇩C by simp
also have "... = D.cod (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using a 1 C⇩U.prj.unit_simps [of a] F⇩U⇩V.unit_simps(5) [of "C⇩U.prj.map⇩0 a"] D⇩V.cod_char⇩S⇩b⇩C
by simp
finally show "D.dom (F (C⇩U.prj.unit a)) = D.cod (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
by blast
qed
have 4: "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.unit a = F (C⇩U.prj.unit a) ⋅⇩D unit (C⇩U.prj.map⇩0 a)"
proof -
have "EQ⇩U⇩VoEQ⇩C'.FH.map⇩0 a = map⇩0 (C⇩U.P⇩0 a)"
using a EQ⇩U⇩VoEQ⇩C'.FH.map⇩0_def D⇩V.src_def C.obj_def F⇩U⇩V.preserves_arr by auto
have seq': "D⇩V.arr (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
proof -
have 5: "D.arr (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using seq by simp
moreover have "src⇩D (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a)) ∈ V"
proof -
have "src⇩D (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))
= map⇩0 (C⇩U.prj.map⇩0 a)"
proof -
have "src⇩D (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))
= src⇩D (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using 5 D.src_vcomp D.vseq_implies_hpar by simp
also have "... = map⇩0 (C⇩U.prj.map⇩0 a)"
using a F⇩U⇩V.unit_simps(1-2) [of "C⇩U.prj.map⇩0 a"] D⇩V.src_def F⇩U⇩V.map⇩0_def
apply simp
using a C⇩U.prj.map⇩0_simps(1) F⇩U⇩V.preserves_arr map⇩0_def by force
finally show ?thesis by blast
qed
moreover have "C.obj (C⇩U.prj.map⇩0 a)"
using a C⇩U.prj.map⇩0_simps(1) C⇩U.obj_char by blast
ultimately show ?thesis
using V_def by blast
qed
moreover have "trg⇩D (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a)) ∈ V"
proof -
have "trg⇩D (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))
= map⇩0 (C⇩U.prj.map⇩0 a)"
proof -
have "trg⇩D (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a)) =
trg⇩D (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using 5 D.trg_vcomp D.vseq_implies_hpar by simp
also have "... = map⇩0 (C⇩U.prj.map⇩0 a)"
using a F⇩U⇩V.unit_simps(1,3) [of "C⇩U.prj.map⇩0 a"] D⇩V.src_def D⇩V.trg_def
F⇩U⇩V.map⇩0_def
apply simp
using a C⇩U.prj.map⇩0_simps(1) F⇩U⇩V.preserves_arr map⇩0_def by force
finally show ?thesis by blast
qed
moreover have "C.obj (C⇩U.prj.map⇩0 a)"
using a C⇩U.prj.map⇩0_simps(1) C⇩U.obj_char by blast
ultimately show ?thesis
using V_def by blast
qed
ultimately show ?thesis
using a D⇩V.arr_char⇩S⇩b⇩C by simp
qed
have "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.unit a =
F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a) ⋅⇩D map⇩0 (C⇩U.P⇩0 a)"
proof -
have 5: "D⇩V.obj (src⇩D (F (C⇩U.P⇩0 a)))"
using a C⇩U.P⇩0_props [of a]
by (metis (no_types, lifting) EQ⇩U⇩VoEQ⇩C'.FH.map⇩0_simps(1)
‹EQ⇩U⇩VoEQ⇩C'.FH.map⇩0 a = map⇩0 (C⇩U.P⇩0 a)› map⇩0_def)
have 6: "D⇩V.emb.unit (map⇩0 (C⇩U.P⇩0 a)) = map⇩0 (C⇩U.P⇩0 a)"
using a 5 D⇩V.emb.unit_char' EQ⇩U⇩VoEQ⇩C'.FH.map⇩0_simps(1)
‹EQ⇩U⇩VoEQ⇩C'.FH.map⇩0 a = map⇩0 (C⇩U.P⇩0 a)› map⇩0_def
by simp
moreover have "D⇩V.emb.unit (map⇩0 (C⇩U.P⇩0 a)) = map⇩0 (C⇩U.prj.map⇩0 a)"
using 6 a C.obj_simps C⇩U.equivalence_data_simps⇩B(7) by simp
moreover have "D⇩V.arr (F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using a C.obj_simps seq' by blast
moreover have "D⇩V.arr (F⇩U⇩V.unit (C⇩U.P⇩0 a))"
using a 1 by simp
moreover have "D⇩V.arr (F (C⇩U.prj.unit a))"
using a F⇩U⇩V.preserves_arr by auto
moreover have "D⇩V.obj (F⇩U⇩V.map⇩0 (C⇩U.P⇩0 a))"
using a 1 F⇩U⇩V.map⇩0_simps(1) by auto
moreover have "F⇩U⇩V.map⇩0 (C⇩U.P⇩0 a) = map⇩0 (C⇩U.P⇩0 a)"
using a ‹EQ⇩U⇩VoEQ⇩C'.FH.map⇩0 a = map⇩0 (C⇩U.P⇩0 a)› map⇩0_def F⇩U⇩V.map⇩0_simps
by auto
ultimately show ?thesis
using a seq EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.unit_char' [of a] EQ⇩U⇩VoEQ⇩C'.FH.unit_char' [of a]
D⇩V.emb.map_def D⇩V.seq_char⇩S⇩b⇩C D⇩V.comp_char D.comp_assoc
by simp
qed
also have "... = F (C⇩U.prj.unit a) ⋅⇩D F⇩U⇩V.unit (C⇩U.prj.map⇩0 a)"
proof -
have "D.dom (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a)) = map⇩0 (C⇩U.P⇩0 a)"
using a 1 F⇩U⇩V.unit_simps(4) [of "C⇩U.prj.map⇩0 a"] D⇩V.dom_char⇩S⇩b⇩C F⇩U⇩V.map⇩0_def
D⇩V.src_def map⇩0_def C.obj_def F⇩U⇩V.preserves_arr
by simp
moreover have "D.arr (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))"
using a ‹D.seq (F (C⇩U.prj.unit a)) (F⇩U⇩V.unit (C⇩U.prj.map⇩0 a))› by blast
ultimately show ?thesis
using a D.comp_arr_dom by auto
qed
also have "... = F (C⇩U.prj.unit a) ⋅⇩D unit (C⇩U.prj.map⇩0 a)"
using a 1 F⇩U⇩V.unit_char' F⇩U.unit_char' by simp
finally show ?thesis by blast
qed
have "(τ⇩0 a ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.unit a) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a] ⋅⇩D 𝗅⇩D[τ⇩0 a] =
(τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a) ⋅⇩D unit (C⇩U.prj.map⇩0 a)) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a] ⋅⇩D 𝗅⇩D[τ⇩0 a]"
using 4 by simp
also have "(τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a) ⋅⇩D unit (C⇩U.prj.map⇩0 a)) ⋅⇩D
𝗋⇩D⇧-⇧1[τ⇩0 a] ⋅⇩D 𝗅⇩D[τ⇩0 a]
= (τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a) ⋅⇩D
unit (C⇩U.prj.map⇩0 a)) ⋅⇩D
(τ⇩0 a ⋆⇩D D.inv (unit (src⇩C (C⇩U.d a)))) ⋅⇩D
D.inv (Φ (C⇩U.d a, src⇩C (C⇩U.d a))) ⋅⇩D
F 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩D
F 𝗅⇩C[C⇩U.d a] ⋅⇩D
Φ (trg⇩C (C⇩U.d a), C⇩U.d a) ⋅⇩D
(unit (trg⇩C (C⇩U.d a)) ⋆⇩D τ⇩0 a)"
proof -
have "𝗅⇩D[τ⇩0 a] = F 𝗅⇩C[C⇩U.d a] ⋅⇩D Φ (trg⇩C (C⇩U.d a), C⇩U.d a) ⋅⇩D
(unit (trg⇩C (C⇩U.d a)) ⋆⇩D τ⇩0 a)"
using a preserves_lunit [of "C⇩U.ε⇩0 a"] C⇩U.counit.ide_map⇩0_obj lunit_coherence
by blast
moreover
have "𝗋⇩D⇧-⇧1[τ⇩0 a] = (τ⇩0 a ⋆⇩D D.inv (unit (src⇩C (C⇩U.d a)))) ⋅⇩D
D.inv (Φ (C⇩U.d a, src⇩C (C⇩U.d a))) ⋅⇩D F 𝗋⇩C⇧-⇧1[C⇩U.d a]"
proof -
have "F 𝗋⇩C⇧-⇧1[C⇩U.d a] =
Φ (C⇩U.d a, src⇩C (C⇩U.d a)) ⋅⇩D (τ⇩0 a ⋆⇩D unit (src⇩C (C⇩U.d a))) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a]"
using preserves_runit(2) [of "C⇩U.ε⇩0 a"] by simp
moreover have 1: "D.iso (Φ (C⇩U.d a, src⇩C (C⇩U.d a)))"
using a C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C C.obj_src [of "C⇩U.d a"]
Φ.components_are_iso [of "(C⇩U.d a, src⇩C (C⇩U.d a))"]
by auto
ultimately have "D.inv (Φ (C⇩U.d a, src⇩C (C⇩U.d a))) ⋅⇩D F 𝗋⇩C⇧-⇧1[C⇩U.d a] =
(τ⇩0 a ⋆⇩D unit (src⇩C (C⇩U.d a))) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a]"
using D.invert_side_of_triangle(1)
[of "F 𝗋⇩C⇧-⇧1[C⇩U.d a]" "Φ (C⇩U.d a, src⇩C (C⇩U.d a))"
"(τ⇩0 a ⋆⇩D unit (src⇩C (C⇩U.d a))) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a]"]
by fastforce
thus ?thesis
using D.invert_side_of_triangle(1)
[of "D.inv (Φ (C⇩U.d a, src⇩C (C⇩U.d a))) ⋅⇩D F 𝗋⇩C⇧-⇧1[C⇩U.d a]"
"τ⇩0 a ⋆⇩D unit (src⇩C (C⇩U.d a))" "𝗋⇩D⇧-⇧1[τ⇩0 a]"]
using C.obj_src [of "C⇩U.d a"] unit_char(2) by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D
((τ⇩0 a ⋆⇩D unit (C⇩U.prj.map⇩0 a)) ⋅⇩D
(τ⇩0 a ⋆⇩D D.inv (unit (C⇩U.prj.map⇩0 a))))) ⋅⇩D
D.inv (Φ (C⇩U.d a, C⇩U.prj.map⇩0 a)) ⋅⇩D
F 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩D F 𝗅⇩C[C⇩U.d a] ⋅⇩D
Φ (a, C⇩U.d a) ⋅⇩D
(unit a ⋆⇩D τ⇩0 a)"
proof -
have "τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a) ⋅⇩D unit (C⇩U.prj.map⇩0 a) =
(τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D (τ⇩0 a ⋆⇩D unit (C⇩U.prj.map⇩0 a))"
using 1 seq D.whisker_left [of "τ⇩0 a"]
by (simp add: F⇩U.unit_char' F⇩U⇩V.unit_char')
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D
D.inv (Φ (C⇩U.d a, C⇩U.prj.map⇩0 a))) ⋅⇩D
F 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩D F 𝗅⇩C[C⇩U.d a] ⋅⇩D
Φ (a, C⇩U.d a) ⋅⇩D
(unit a ⋆⇩D τ⇩0 a)"
proof -
have "(τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D
((τ⇩0 a ⋆⇩D unit (C⇩U.prj.map⇩0 a)) ⋅⇩D
(τ⇩0 a ⋆⇩D D.inv (unit (C⇩U.prj.map⇩0 a))))
= (τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D
(τ⇩0 a ⋆⇩D unit (C⇩U.prj.map⇩0 a) ⋅⇩D
D.inv (unit (C⇩U.prj.map⇩0 a)))"
proof -
have "D.seq (unit (C⇩U.prj.map⇩0 a)) (D.inv (unit (C⇩U.prj.map⇩0 a)))"
using a 1 unit_char(1-2) C⇩U.obj_char
by (intro D.seqI) auto
thus ?thesis
using a D.whisker_left [of "τ⇩0 a"] by simp
qed
also have "... = (τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D (τ⇩0 a ⋆⇩D F (C⇩U.prj.map⇩0 a))"
using a 1 unit_char(1-2) C⇩U.obj_char unit_simps [of "C⇩U.prj.map⇩0 a"]
by (simp add: D.comp_arr_inv')
also have "... = τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a) ⋅⇩D F (C⇩U.prj.map⇩0 a)"
using seq D.whisker_left by auto
also have "... = τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)"
using D.comp_arr_dom by simp
finally have "(τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D
((τ⇩0 a ⋆⇩D unit (C⇩U.prj.map⇩0 a)) ⋅⇩D
(τ⇩0 a ⋆⇩D D.inv (unit (C⇩U.prj.map⇩0 a))))
= τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = D.inv (Φ (C⇩U.d a, C⇩U.P a)) ⋅⇩D
(F (C⇩U.d a ⋆⇩C C⇩U.prj.unit a) ⋅⇩D
F 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩D F 𝗅⇩C[C⇩U.d a]) ⋅⇩D
Φ (a, C⇩U.d a) ⋅⇩D
(unit a ⋆⇩D τ⇩0 a)"
proof -
have "(τ⇩0 a ⋆⇩D F (C⇩U.prj.unit a)) ⋅⇩D D.inv (Φ (C⇩U.d a, C⇩U.prj.map⇩0 a)) =
D.inv (Φ (C⇩U.d a, C⇩U.P a)) ⋅⇩D F (C⇩U.d a ⋆⇩C C⇩U.prj.unit a)"
proof -
have "C.VV.arr (C⇩U.d a, C⇩U.prj.unit a)"
using a C.VV.arr_char⇩S⇩b⇩C by simp
thus ?thesis
using a Φ.inv_naturality [of "(C⇩U.d a, C⇩U.prj.unit a)"] C⇩U.prj.unit_simps [of a]
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (D.inv (Φ (C⇩U.d a, C⇩U.P a)) ⋅⇩D
F (C⇩U.counit⇩1 a ⋅⇩C (a ⋆⇩C C⇩U.d a)) ⋅⇩D
Φ (a, C⇩U.d a)) ⋅⇩D
(unit a ⋆⇩D τ⇩0 a)"
proof -
have "F (C⇩U.d a ⋆⇩C C⇩U.prj.unit a) ⋅⇩D F 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩D F 𝗅⇩C[C⇩U.d a] =
F ((C⇩U.d a ⋆⇩C C⇩U.prj.unit a) ⋅⇩C 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩C 𝗅⇩C[C⇩U.d a])"
using a by simp
also have "... = F (C⇩U.counit⇩1 a ⋅⇩C (a ⋆⇩C C⇩U.d a))"
proof -
have "C⇩U.EoP.unit a = C⇩U.prj.unit a"
using a 1 C⇩U.emb.map_def C⇩U.EoP.unit_char' C⇩U.emb.unit_char' C.comp_arr_dom
by simp
thus ?thesis
using a C⇩U.counit.respects_unit [of a] C⇩U.I⇩C.unit_char' [of a] by simp
qed
finally have "F (C⇩U.d a ⋆⇩C C⇩U.prj.unit a) ⋅⇩D F 𝗋⇩C⇧-⇧1[C⇩U.d a] ⋅⇩D F 𝗅⇩C[C⇩U.d a] =
F (C⇩U.counit⇩1 a ⋅⇩C (a ⋆⇩C C⇩U.d a))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = τ⇩1 a ⋅⇩D (unit a ⋆⇩D τ⇩0 a)"
proof -
have "D.inv (Φ (C⇩U.d a, C⇩U.P a)) ⋅⇩D F (C⇩U.ε⇩1 a ⋅⇩C (a ⋆⇩C C⇩U.d a)) ⋅⇩D Φ (a, C⇩U.d a)
= D.inv (Φ (C⇩U.d a, C⇩U.P a)) ⋅⇩D F (C⇩U.ε⇩1 a) ⋅⇩D Φ (a, C⇩U.d a)"
using a C.obj_def' C.comp_arr_dom
by (metis C⇩U.counit⇩1_simps(1) C⇩U.counit⇩1_simps(4) C.objE)
also have "... = τ⇩1 a"
using a by auto
finally have "D.inv (Φ (C⇩U.d a, C⇩U.P a)) ⋅⇩D F (C⇩U.ε⇩1 a ⋅⇩C (a ⋆⇩C C⇩U.d a)) ⋅⇩D Φ (a, C⇩U.d a)
= τ⇩1 a"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
finally show "(τ⇩0 a ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.unit a) ⋅⇩D 𝗋⇩D⇧-⇧1[τ⇩0 a] ⋅⇩D 𝗅⇩D[τ⇩0 a]
= τ⇩1 a ⋅⇩D (unit a ⋆⇩D τ⇩0 a)"
by blast
qed
show "⋀f g. ⟦C.ide f; C.ide g; src⇩C g = trg⇩C f⟧ ⟹
(τ⇩0 (trg⇩C g) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp (g, f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map g,
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f] ⋅⇩D
(τ⇩1 g ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (src⇩C g), EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D 𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]
= τ⇩1 (g ⋆⇩C f) ⋅⇩D (Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f))"
proof -
fix f g
assume f: "C.ide f" and g: "C.ide g" and fg: "src⇩C g = trg⇩C f"
have 1: "C.ide f ∧ C.ide g ∧ C.ide (C⇩U.P f) ∧ C.ide (C⇩U.P g) ∧
C.ide (C⇩U.d (trg⇩C f)) ∧ C.ide (C⇩U.d (trg⇩C g)) ∧
src⇩C (C⇩U.d (trg⇩C f)) = trg⇩C (C⇩U.P f) ∧
src⇩C (C⇩U.d (trg⇩C g)) = trg⇩C (C⇩U.P g) ∧
trg⇩C (C⇩U.d (trg⇩C f)) = src⇩C g"
using f g fg C⇩U.emb.map⇩0_def C⇩U.emb.map_def C⇩U.obj_char C⇩U.P⇩0_props(1) C.obj_simps(2)
C⇩U.prj.preserves_ide C⇩U.ide_char⇩S⇩b⇩C
by auto
have "τ⇩1 (g ⋆⇩C f) ⋅⇩D (Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f))
= (D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
F ((C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.EoP.cmp (g, f)) ⋅⇩C
𝖺⇩C[C⇩U.d (trg⇩C g), C⇩U.P g, C⇩U.P f] ⋅⇩C
(C⇩U.ε⇩1 g ⋆⇩C C⇩U.P f) ⋅⇩C
C.inv 𝖺⇩C[g, C⇩U.d (src⇩C g), C⇩U.P f] ⋅⇩C
(g ⋆⇩C C⇩U.ε⇩1 f) ⋅⇩C
𝖺⇩C[g, f, C⇩U.d (src⇩C f)]) ⋅⇩D
Φ (g ⋆⇩C f, C⇩U.d (src⇩C f))) ⋅⇩D
(Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f))"
using f g fg C⇩U.emb.map_def C⇩U.counit.respects_hcomp [of f g] D.comp_arr_dom
by simp
also have "... = ((D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f)))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
(τ⇩1 g ⋆⇩D F (C⇩U.P f)) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D
(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D
(D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f))) ⋅⇩D
Φ (g ⋆⇩C f, C⇩U.d (src⇩C f))) ⋅⇩D
(Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f))"
proof -
have "F ((C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.EoP.cmp (g, f)) ⋅⇩C
𝖺⇩C[C⇩U.d (trg⇩C g), C⇩U.P g, C⇩U.P f] ⋅⇩C
(C⇩U.ε⇩1 g ⋆⇩C C⇩U.P f) ⋅⇩C
C.inv 𝖺⇩C[g, C⇩U.d (src⇩C g), C⇩U.P f] ⋅⇩C
(g ⋆⇩C C⇩U.ε⇩1 f) ⋅⇩C
𝖺⇩C[g, f, C⇩U.d (src⇩C f)])
= F (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.EoP.cmp (g, f)) ⋅⇩D
F 𝖺⇩C[C⇩U.d (trg⇩C g), C⇩U.P g, C⇩U.P f] ⋅⇩D
F (C⇩U.ε⇩1 g ⋆⇩C C⇩U.P f) ⋅⇩D
F (C.inv 𝖺⇩C[g, C⇩U.d (src⇩C g), C⇩U.P f]) ⋅⇩D
F (g ⋆⇩C C⇩U.ε⇩1 f) ⋅⇩D
F 𝖺⇩C[g, f, C⇩U.d (src⇩C f)]"
proof -
have "C.arr ((C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.EoP.cmp (g, f)) ⋅⇩C
𝖺⇩C[C⇩U.d (trg⇩C g), C⇩U.P g, C⇩U.P f] ⋅⇩C
(C⇩U.ε⇩1 g ⋆⇩C C⇩U.P f) ⋅⇩C
C.inv 𝖺⇩C[g, C⇩U.d (src⇩C g), C⇩U.P f] ⋅⇩C
(g ⋆⇩C C⇩U.ε⇩1 f) ⋅⇩C
𝖺⇩C[g, f, C⇩U.d (src⇩C f)])"
using f g fg 1 C⇩U.emb.map_def C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C⇩U.EoP.FF_def
by (intro C.seqI C.hseqI') auto
thus ?thesis
using f g fg by fastforce
qed
also have "... = Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f)) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))) ⋅⇩D
((D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P g ⋆⇩C C⇩U.P f)) ⋅⇩D
Φ (C⇩U.d (trg⇩C g), C⇩U.P g ⋆⇩C C⇩U.P f)) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f))) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
(D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P g)) ⋆⇩D F (C⇩U.P f)) ⋅⇩D
((D.inv (Φ (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g, C⇩U.P f)) ⋅⇩D
Φ (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g, C⇩U.P f)) ⋅⇩D
(F (C⇩U.ε⇩1 g) ⋆⇩D F (C⇩U.P f))) ⋅⇩D
((D.inv (Φ (g ⋆⇩C C⇩U.d (trg⇩C f), C⇩U.P f)) ⋅⇩D
Φ (g ⋆⇩C C⇩U.d (trg⇩C f), C⇩U.P f)) ⋅⇩D
(Φ (g, C⇩U.d (trg⇩C f)) ⋆⇩D F (C⇩U.P f))) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
(F g ⋆⇩D D.inv (Φ (C⇩U.d (trg⇩C f), C⇩U.P f))) ⋅⇩D
((D.inv (Φ (g, C⇩U.d (trg⇩C f) ⋆⇩C C⇩U.P f)) ⋅⇩D
Φ (g, C⇩U.d (trg⇩C f) ⋆⇩C C⇩U.P f)) ⋅⇩D
(F g ⋆⇩D F (C⇩U.ε⇩1 f))) ⋅⇩D
((D.inv (Φ (g, f ⋆⇩C C⇩U.d (src⇩C f))) ⋅⇩D
Φ (g, f ⋆⇩C C⇩U.d (src⇩C f))) ⋅⇩D
(F g ⋆⇩D Φ (f, C⇩U.d (src⇩C f)))) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D
(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D
D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f)))"
using 1 f g fg preserves_hcomp preserves_assoc C⇩U.emb.map_def
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def D.comp_assoc
by simp
also have "... = Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f)) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
((D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P g)) ⋆⇩D F (C⇩U.P f)) ⋅⇩D
(F (C⇩U.ε⇩1 g) ⋆⇩D F (C⇩U.P f)) ⋅⇩D
(Φ (g, C⇩U.d (trg⇩C f)) ⋆⇩D F (C⇩U.P f))) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
((F g ⋆⇩D D.inv (Φ (C⇩U.d (trg⇩C f), C⇩U.P f))) ⋅⇩D
(F g ⋆⇩D F (C⇩U.ε⇩1 f)) ⋅⇩D
(F g ⋆⇩D Φ (f, C⇩U.d (src⇩C f)))) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D
(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D
D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f)))"
proof -
have "(D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P g ⋆⇩C C⇩U.P f)) ⋅⇩D
Φ (C⇩U.d (trg⇩C g), C⇩U.P g ⋆⇩C C⇩U.P f)) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f))
= τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f)"
using f g fg 1 C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
moreover have "(D.inv (Φ (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g, C⇩U.P f)) ⋅⇩D
Φ (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g, C⇩U.P f)) ⋅⇩D
(F (C⇩U.ε⇩1 g) ⋆⇩D F (C⇩U.P f))
= F (C⇩U.ε⇩1 g) ⋆⇩D F (C⇩U.P f)"
proof -
have "(D.inv (Φ (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g, C⇩U.P f)) ⋅⇩D
Φ (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g, C⇩U.P f)) ⋅⇩D
(F (C⇩U.ε⇩1 g) ⋆⇩D F (C⇩U.P f))
= (F (C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.P g) ⋆⇩D F (C⇩U.P f)) ⋅⇩D
(F (C⇩U.ε⇩1 g) ⋆⇩D F (C⇩U.P f))"
using f g fg 1 C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
also have "... = F (C⇩U.counit⇩1 g) ⋆⇩D F (C⇩U.P f)"
using g D.comp_cod_arr f g fg C⇩U.P⇩0_props(1) C⇩U.emb.map_def by simp
finally show ?thesis by blast
qed
moreover have "(D.inv (Φ (g ⋆⇩C C⇩U.d (trg⇩C f), C⇩U.P f)) ⋅⇩D
(Φ (g ⋆⇩C C⇩U.d (trg⇩C f), C⇩U.P f))) ⋅⇩D
(Φ (g, C⇩U.d (trg⇩C f)) ⋆⇩D F (C⇩U.P f))
= (Φ (g, C⇩U.d (trg⇩C f)) ⋆⇩D F (C⇩U.P f))"
using f g fg 1 C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
moreover have "(D.inv (Φ (g, C⇩U.d (trg⇩C f) ⋆⇩C C⇩U.P f)) ⋅⇩D
Φ (g, C⇩U.d (trg⇩C f) ⋆⇩C C⇩U.P f)) ⋅⇩D
(F g ⋆⇩D F (C⇩U.counit⇩1 f))
= F g ⋆⇩D F (C⇩U.counit⇩1 f)"
using f g fg 1 C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
moreover have "(D.inv (Φ (g, f ⋆⇩C C⇩U.d (src⇩C f))) ⋅⇩D
Φ (g, f ⋆⇩C C⇩U.d (src⇩C f))) ⋅⇩D
(F g ⋆⇩D Φ (f, C⇩U.d (src⇩C f)))
= F g ⋆⇩D Φ (f, C⇩U.d (src⇩C f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f)) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
(τ⇩1 g ⋆⇩D F (C⇩U.P f)) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D
(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D
D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f)))"
proof -
have "(F g ⋆⇩D D.inv (Φ (C⇩U.d (trg⇩C f), C⇩U.P f))) ⋅⇩D
(F g ⋆⇩D F (C⇩U.counit⇩1 f)) ⋅⇩D
(F g ⋆⇩D Φ (f, C⇩U.d (src⇩C f)))
= F g ⋆⇩D τ⇩1 f"
using f g fg 1 D.whisker_left C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
by simp
moreover have "(D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P g)) ⋆⇩D F (C⇩U.P f)) ⋅⇩D
(F (C⇩U.counit⇩1 g) ⋆⇩D F (C⇩U.P f)) ⋅⇩D
(Φ (g, C⇩U.d (trg⇩C f)) ⋆⇩D F (C⇩U.P f))
= τ⇩1 g ⋆⇩D F (C⇩U.P f)"
using f g fg 1 D.whisker_right C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
C.VV.ide_char⇩S⇩b⇩C Φ.components_are_iso C⇩U.emb.map_def
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
finally have "F ((C⇩U.d (trg⇩C g) ⋆⇩C C⇩U.EoP.cmp (g, f)) ⋅⇩C
𝖺⇩C[C⇩U.d (trg⇩C g), C⇩U.P g, C⇩U.P f] ⋅⇩C
(C⇩U.ε⇩1 g ⋆⇩C C⇩U.P f) ⋅⇩C
C.inv 𝖺⇩C[g, C⇩U.d (src⇩C g), C⇩U.P f] ⋅⇩C
(g ⋆⇩C C⇩U.ε⇩1 f) ⋅⇩C
𝖺⇩C[g, f, C⇩U.d (src⇩C f)])
= Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f)) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
(τ⇩1 g ⋆⇩D F (C⇩U.P f)) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D
(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D
D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f)))"
by blast
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = ((τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f))) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
(τ⇩1 g ⋆⇩D F (C⇩U.P f)) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]"
proof -
have "(D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f)))
= τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))"
proof -
have "(D.inv (Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f)))
= (τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.P (g ⋆⇩C f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f)))"
proof -
have "D.iso (Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f)))"
using f g fg Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C C.VV.arr_char⇩S⇩b⇩C
by (metis (no_types, lifting) C⇩U.prj.preserves_ide C⇩U.P_simps⇩B(3) C⇩U.ide_char⇩S⇩b⇩C
C⇩U.counit.ide_map⇩0_obj C⇩U.equivalence_data_simps⇩B(7) C.hcomp_simps(2)
C.ideD(1) C.ide_hcomp C.obj_trg cmp_components_are_iso)
moreover have "D.dom (Φ (C⇩U.d (trg⇩C g), C⇩U.P (g ⋆⇩C f))) =
τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.P (g ⋆⇩C f))"
using f g fg C⇩U.ide_char⇩S⇩b⇩C C⇩U.equivalence_data_simps⇩B(7) C⇩U.prj.preserves_ide
cmp_simps(4)
by simp
ultimately show ?thesis
using D.comp_inv_arr' by auto
qed
also have "... = τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.EoP.cmp (g, f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.whisker_left [of "τ⇩0 (trg⇩C g)"]
by simp
also have "... = τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_cod_arr
by simp
finally show ?thesis by blast
qed
moreover have "𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D
(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D
(D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f))) ⋅⇩D
Φ (g ⋆⇩C f, C⇩U.d (src⇩C f))) ⋅⇩D
(Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f))
= 𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]"
proof -
have "D.inv (Φ (g ⋆⇩C f, C⇩U.d (src⇩C f))) ⋅⇩D Φ (g ⋆⇩C f, C⇩U.d (src⇩C f)) =
F (g ⋆⇩C f) ⋆⇩D F (C⇩U.d (src⇩C f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
moreover have "(F (g ⋆⇩C f) ⋆⇩D F (C⇩U.d (src⇩C f))) ⋅⇩D (Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f)) =
Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f)"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
by simp
moreover have "(D.inv (Φ (g, f)) ⋆⇩D τ⇩0 (src⇩C f)) ⋅⇩D (Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f)) =
(F g ⋆⇩D F f) ⋆⇩D τ⇩0 (src⇩C f)"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
Φ.components_are_iso C.VV.ide_char⇩S⇩b⇩C FF_def
D.whisker_right [of "τ⇩0 (src⇩C f)" "D.inv (Φ (g, f))" "Φ (g, f)"]
by simp
moreover have "𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)] ⋅⇩D ((F g ⋆⇩D F f) ⋆⇩D τ⇩0 (src⇩C f)) =
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩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 "... = (τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f)) ⋅⇩D
Φ (C⇩U.P g, C⇩U.P f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), F (C⇩U.P g), F (C⇩U.P f)] ⋅⇩D
(τ⇩1 g ⋆⇩D F (C⇩U.P f)) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (trg⇩C f), F (C⇩U.P f)] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]"
proof -
have "(τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f))) ⋅⇩D
(τ⇩0 (trg⇩C g) ⋆⇩D Φ (C⇩U.P g, C⇩U.P f))
= τ⇩0 (trg⇩C g) ⋆⇩D F (C⇩U.EoP.cmp (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f)"
proof -
have "D.arr (F (C⇩U.EoP.cmp (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C C⇩U.EoP.FF_def
C⇩U.emb.map_def
by (intro D.seqI) auto
thus ?thesis
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C
D.whisker_left [of "τ⇩0 (trg⇩C g)"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (τ⇩0 (trg⇩C g) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp (g, f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map g,
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f] ⋅⇩D
(τ⇩1 g ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (src⇩C g), EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D
𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]"
proof -
have "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp (g, f) =
F (C⇩U.EoP.cmp (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f)"
proof -
have "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp (g, f) =
F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D ((F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.Φ⇩P (g, f))) ⋅⇩D
Φ (C⇩U.P g, C⇩U.P f)) ⋅⇩D D⇩V.Φ⇩E (F (C⇩U.P g), F (C⇩U.P f))"
proof -
have 3: "D.arr (F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.Φ⇩P (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def C⇩U.prj.FF_def
C⇩U.Φ⇩P_in_hom [of g f] C⇩U.hcomp_def C⇩U.in_hom_char⇩S⇩b⇩C
by (intro D.seqI) auto
have 4: "D⇩V.arr (Φ (C⇩U.P g, C⇩U.P f))"
using f g fg 1 cmp_in_hom(1) [of "C⇩U.P g" "C⇩U.P f"] C⇩U.prj.map⇩0_simps(1)
C⇩U.obj_char D⇩V.arr_char⇩S⇩b⇩C V_def
by auto
moreover have 5: "D⇩V.arr (F (C⇩U.Φ⇩P (g, f)))"
using f g fg 1 C⇩U.Φ⇩P_simps [of g f] F⇩U⇩V.preserves_arr by presburger
moreover have 6: "D⇩V.arr (F (C⇩U.P (g ⋆⇩C f)))"
using f g fg C⇩U.P_simps(1) [of "g ⋆⇩C f"] F⇩U⇩V.preserves_arr by simp
moreover have 7: "D⇩V.arr (F (C⇩U.Φ⇩P (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f))"
using 3 4 5 D⇩V.seq_char⇩S⇩b⇩C [of "F (C⇩U.Φ⇩P (g, f))" "Φ (C⇩U.P g, C⇩U.P f)"]
D⇩V.comp_char [of "F (C⇩U.Φ⇩P (g, f))" "Φ (C⇩U.P g, C⇩U.P f)"]
by auto
moreover have "D.seq (F (C⇩U.Φ⇩P (g, f))) (Φ (C⇩U.P g, C⇩U.P f))"
using 3 by blast
moreover have "D.seq (F (C⇩U.P (g ⋆⇩C f)))
(F (C⇩U.Φ⇩P (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f))"
using 3 by blast
moreover have "D⇩V.arr (F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.Φ⇩P (g, f)) ⋅⇩D
Φ (C⇩U.P g, C⇩U.P f))"
using f g fg 3 6 7 D.vseq_implies_hpar V_def
by (metis (no_types, lifting) D⇩V.comp_char D⇩V.seq_char⇩S⇩b⇩C)
ultimately show ?thesis
using f g fg EQ⇩DoEQ⇩U⇩VoEQ⇩C'.FH.cmp_def EQ⇩U⇩VoEQ⇩C'.FH.cmp_def
C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C⇩U.VV.arr_char⇩S⇩b⇩C D⇩V.comp_char
D⇩V.emb.map_def D.comp_assoc
by simp
qed
also have "... = F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.Φ⇩P (g, f)) ⋅⇩D
Φ (C⇩U.P g, C⇩U.P f) ⋅⇩D D⇩V.Φ⇩E (F (C⇩U.P g), F (C⇩U.P f))"
proof -
have "F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.Φ⇩P (g, f)) = F (C⇩U.Φ⇩P (g, f))"
proof -
have 1: "C.arr (C⇩U.Φ⇩P (g, f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C
by (metis (no_types, lifting) C⇩U.prj.preserves_hcomp C⇩U.prj.preserves_ide
C⇩U.ide_char⇩S⇩b⇩C C⇩U.seq_char⇩S⇩b⇩C C.ideD(1) C.ideD(3) C.ide_hcomp C.seqE)
moreover have "C.cod (C⇩U.Φ⇩P (g, f)) = C⇩U.P (g ⋆⇩C f)"
using f g fg C.VV.arr_char⇩S⇩b⇩C C⇩U.Φ⇩P_in_hom(2) [of g f]
C⇩U.hcomp_def C⇩U.in_hom_char⇩S⇩b⇩C by auto
ultimately have "D.cod (F (C⇩U.Φ⇩P (g, f))) = F (C⇩U.P (g ⋆⇩C f))"
using f g fg C.VV.arr_char⇩S⇩b⇩C preserves_cod [of "C⇩U.Φ⇩P (g, f)"]
by simp
thus ?thesis
using 1 f g fg D.comp_cod_arr [of "F (C⇩U.Φ⇩P (g, f))" "F (C⇩U.P (g ⋆⇩C f))"]
by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = F (C⇩U.P (g ⋆⇩C f)) ⋅⇩D F (C⇩U.Φ⇩P (g, f)) ⋅⇩D
F (C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f)"
proof -
have "D⇩V.Φ⇩E (F (C⇩U.P g), F (C⇩U.P f)) = F (C⇩U.P g) ⋆⇩D F (C⇩U.P f)"
using f g fg D⇩V.emb.cmp_def D⇩V.VV.arr_char⇩S⇩b⇩C D⇩V.src_def D⇩V.trg_def
F⇩U⇩V.preserves_arr
by auto
moreover have "C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f) = C⇩U.P g ⋆⇩C C⇩U.P f"
using f g fg C⇩U.VV.arr_char⇩S⇩b⇩C C⇩U.emb.cmp_def by simp
ultimately
have "Φ (C⇩U.P g, C⇩U.P f) ⋅⇩D D⇩V.Φ⇩E (F (C⇩U.P g), F (C⇩U.P f)) =
F (C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f)"
using f g fg C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C.VV.cod_char⇩S⇩b⇩C FF_def
Φ.naturality [of "(C⇩U.P g, C⇩U.P f)"]
by simp
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = F (C⇩U.P (g ⋆⇩C f) ⋅⇩C C⇩U.Φ⇩P (g, f) ⋅⇩C C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f)) ⋅⇩D
Φ (C⇩U.P g, C⇩U.P f)"
proof -
have "C.arr (C⇩U.P (g ⋆⇩C f) ⋅⇩C C⇩U.Φ⇩P (g, f) ⋅⇩C C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
proof (intro C.seqI)
show "C.arr (C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
using f g fg by auto
show 1: "C.arr (C⇩U.Φ⇩P (g, f))"
using f g fg C⇩U.arr_char⇩S⇩b⇩C C⇩U.Φ⇩P_simps(1) [of g f] by auto
show "C.dom (C⇩U.Φ⇩P (g, f)) = C.cod (C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
proof -
have "C.dom (C⇩U.Φ⇩P (g, f)) = C⇩U.hcomp (C⇩U.P g) (C⇩U.P f)"
using f g fg C⇩U.VV.arr_char⇩S⇩b⇩C C⇩U.Φ⇩P_simps(4) [of g f]
C⇩U.dom_char⇩S⇩b⇩C [of "C⇩U.Φ⇩P (g, f)"]
by fastforce
also have "... = C⇩U.P g ⋆⇩C (C⇩U.P f)"
using f g fg by auto
also have "... = C.cod (C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
using f g fg C⇩U.emb.cmp_def [of "(C⇩U.P g, C⇩U.P f)"] by auto
finally show ?thesis by blast
qed
show "C.arr (C⇩U.P (g ⋆⇩C f))"
using f g fg by simp
show "C.dom (C⇩U.P (g ⋆⇩C f)) =
C.cod (C⇩U.Φ⇩P (g, f) ⋅⇩C C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
proof -
have "C.dom (C⇩U.P (g ⋆⇩C f)) = C⇩U.P (g ⋆⇩C f)"
using f g fg by simp
also have "... = C.cod (C⇩U.Φ⇩P (g, f))"
using f g fg C⇩U.Φ⇩P_simps(5) [of g f] C⇩U.cod_char⇩S⇩b⇩C [of "C⇩U.Φ⇩P (g, f)"]
by fastforce
also have "... = C.cod (C⇩U.Φ⇩P (g, f) ⋅⇩C C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
proof -
have 2: "C⇩U.hcomp (C⇩U.P g) (C⇩U.P f) = C⇩U.P g ⋆⇩C C⇩U.P f"
using f g fg by auto
have "C⇩U.arr (C⇩U.P g ⋆⇩C C⇩U.P f)"
proof -
have "C⇩U.arr (C⇩U.hcomp (C⇩U.P g) (C⇩U.P f))"
using f g fg by simp
thus ?thesis
using 2 by simp
qed
moreover have "C.dom (C⇩U.Φ⇩P (g, f)) = C⇩U.P g ⋆⇩C C⇩U.P f"
using f g fg 2 C⇩U.VV.arr_char⇩S⇩b⇩C C⇩U.Φ⇩P_simps(4) [of g f]
C⇩U.dom_char⇩S⇩b⇩C [of "C⇩U.Φ⇩P (g, f)"]
by fastforce
ultimately have "C.arr (C⇩U.Φ⇩P (g, f) ⋅⇩C C⇩U.Φ⇩E (C⇩U.P g, C⇩U.P f))"
using f g fg 1 C⇩U.VV.arr_char⇩S⇩b⇩C C⇩U.VV.cod_char⇩S⇩b⇩C C⇩U.hcomp_char
C⇩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⇩U.EoP.cmp (g, f)) ⋅⇩D Φ (C⇩U.P g, C⇩U.P f)"
using f g fg C⇩U.EoP.cmp_def C.VV.arr_char⇩S⇩b⇩C C.VV.dom_char⇩S⇩b⇩C C⇩U.emb.map_def
by simp
finally show ?thesis by blast
qed
moreover have "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map g = F (C⇩U.P g)"
proof -
have "D⇩V.arr (F (C⇩U.P g))"
using g D⇩V.arr_char⇩S⇩b⇩C C⇩U.P_simps(1) C.ideD(1) F⇩U⇩V.preserves_arr by presburger
thus ?thesis
using g D⇩V.emb.map_def by simp
qed
moreover have "EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f = F (C⇩U.P f)"
proof -
have "D⇩V.arr (F (C⇩U.P f))"
using f D⇩V.arr_char⇩S⇩b⇩C C⇩U.P_simps(1) C.ideD(1) F⇩U⇩V.preserves_arr by presburger
thus ?thesis
using f D⇩V.emb.map_def by simp
qed
ultimately show ?thesis
using fg by argo
qed
finally show "(τ⇩0 (trg⇩C g) ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp (g, f)) ⋅⇩D
𝖺⇩D[τ⇩0 (trg⇩C g), EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map g,
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f] ⋅⇩D
(τ⇩1 g ⋆⇩D EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f) ⋅⇩D
D.inv 𝖺⇩D[F g, τ⇩0 (src⇩C g), EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map f] ⋅⇩D
(F g ⋆⇩D τ⇩1 f) ⋅⇩D 𝖺⇩D[F g, F f, τ⇩0 (src⇩C f)]
= τ⇩1 (g ⋆⇩C f) ⋅⇩D (Φ (g, f) ⋆⇩D τ⇩0 (src⇩C f))"
by argo
qed
qed
interpretation EQ: equivalence_of_bicategories_and_pseudonatural_equivalence_left
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩1
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩1
F Φ τ⇩0 τ⇩1
proof -
interpret E: equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩1
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩1
using induces_equivalence_of_bicategories by blast
interpret τ: pseudonatural_equivalence
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp F Φ τ⇩0 τ⇩1
..
show "equivalence_of_bicategories_and_pseudonatural_equivalence_left
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.left_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_cmp
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.unit⇩1
EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩0 EQ⇩DoEQ⇩U⇩VoEQ⇩C'.counit⇩1
F Φ τ⇩0 τ⇩1"
..
qed
definition right_map
where "right_map ≡ EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_map"
definition right_cmp
where "right_cmp ≡ EQ⇩DoEQ⇩U⇩VoEQ⇩C'.right_cmp"
definition unit⇩0
where "unit⇩0 ≡ EQ.unit.map⇩0"
definition unit⇩1
where "unit⇩1 ≡ EQ.unit.map⇩1"
definition counit⇩0
where "counit⇩0 ≡ EQ.counit.map⇩0"
definition counit⇩1
where "counit⇩1 ≡ EQ.counit.map⇩1"
theorem extends_to_equivalence_of_bicategories:
shows "equivalence_of_bicategories V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F Φ right_map right_cmp unit⇩0 unit⇩1 counit⇩0 counit⇩1"
unfolding right_map_def right_cmp_def unit⇩0_def unit⇩1_def counit⇩0_def counit⇩1_def
..
end
locale converse_equivalence_pseudofunctor =
C: bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C +
D: bicategory V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D +
F: equivalence_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F
for V⇩C :: "'c comp" (infixr ‹⋅⇩C› 55)
and H⇩C :: "'c comp" (infixr ‹⋆⇩C› 53)
and 𝖺⇩C :: "'c ⇒ 'c ⇒ 'c ⇒ 'c" (‹𝖺⇩C[_, _, _]›)
and 𝗂⇩C :: "'c ⇒ 'c" (‹𝗂⇩C[_]›)
and src⇩C :: "'c ⇒ 'c"
and trg⇩C :: "'c ⇒ 'c"
and V⇩D :: "'d comp" (infixr ‹⋅⇩D› 55)
and H⇩D :: "'d comp" (infixr ‹⋆⇩D› 53)
and 𝖺⇩D :: "'d ⇒ 'd ⇒ 'd ⇒ 'd" (‹𝖺⇩D[_, _, _]›)
and 𝗂⇩D :: "'d ⇒ 'd" (‹𝗂⇩D[_]›)
and src⇩D :: "'d ⇒ 'd"
and trg⇩D :: "'d ⇒ 'd"
and F :: "'c ⇒ 'd"
and Φ⇩F :: "'c * 'c ⇒ 'd"
begin
interpretation E: equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F Φ⇩F F.right_map F.right_cmp F.unit⇩0 F.unit⇩1 F.counit⇩0 F.counit⇩1
using F.extends_to_equivalence_of_bicategories by simp
interpretation E': converse_equivalence_of_bicategories
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F Φ⇩F F.right_map F.right_cmp F.unit⇩0 F.unit⇩1 F.counit⇩0 F.counit⇩1
..
sublocale equivalence_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F.right_map F.right_cmp
using E'.equivalence_pseudofunctor_left by simp
lemma is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
F.right_map F.right_cmp"
..
end
definition equivalent_bicategories
where "equivalent_bicategories V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C ≡
∃F Φ. equivalence_pseudofunctor
V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ"
lemma equivalent_bicategories_reflexive:
assumes "bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C"
shows "equivalent_bicategories V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C"
proof -
interpret bicategory V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C
using assms by simp
interpret I: identity_pseudofunctor V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C ..
show ?thesis
using I.equivalence_pseudofunctor_axioms equivalent_bicategories_def by blast
qed
lemma equivalent_bicategories_symmetric:
assumes "equivalent_bicategories V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D"
shows "equivalent_bicategories V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C"
proof -
obtain F Φ⇩F where F: "equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F"
using assms equivalent_bicategories_def by blast
interpret F: equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F
using F by simp
interpret G: converse_equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F
..
show ?thesis
using G.is_equivalence_pseudofunctor equivalent_bicategories_def by blast
qed
lemma equivalent_bicategories_transitive:
assumes "equivalent_bicategories V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C"
and "equivalent_bicategories V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D"
shows "equivalent_bicategories V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D"
proof -
obtain F Φ⇩F where F: "equivalence_pseudofunctor
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F"
using assms(1) equivalent_bicategories_def by blast
obtain G Φ⇩G where G: "equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G"
using assms(2) equivalent_bicategories_def by blast
interpret F: equivalence_pseudofunctor
V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C F Φ⇩F
using F by simp
interpret G: equivalence_pseudofunctor
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D G Φ⇩G
using G by simp
interpret GF: composite_equivalence_pseudofunctor V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B
V⇩C H⇩C 𝖺⇩C 𝗂⇩C src⇩C trg⇩C V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D F Φ⇩F G Φ⇩G ..
show "equivalent_bicategories V⇩B H⇩B 𝖺⇩B 𝗂⇩B src⇩B trg⇩B V⇩D H⇩D 𝖺⇩D 𝗂⇩D src⇩D trg⇩D"
using equivalent_bicategories_def GF.equivalence_pseudofunctor_axioms by blast
qed
end