open import Cat.Functor.Adjoint.Compose
open import Cat.Functor.Naturality
open import Cat.Functor.Properties
open import Cat.Functor.Adjoint
open import Cat.Functor.Base
open import Cat.Prelude
import Cat.Functor.Reasoning as Fr
import Cat.Reasoning
module Cat.Functor.Equivalence where
Equivalences🔗
A functor is an equivalence of categories when it has a right adjoint , with the unit and counit natural transformations being natural isomorphisms. This immediately implies that our adjoint pair extends to an adjoint triple .
record is-equivalence (F : Functor C D) : Type (adj-level C D) where
private
module C = Cat.Reasoning C
module D = Cat.Reasoning D
module [C,C] = Cat.Reasoning Cat[ C , C ]
module [D,D] = Cat.Reasoning Cat[ D , D ]
field
: Functor D C
F⁻¹ : F ⊣ F⁻¹
F⊣F⁻¹
open _⊣_ F⊣F⁻¹ public
field
: ∀ x → C.is-invertible (unit.η x)
unit-iso : ∀ x → D.is-invertible (counit.ε x) counit-iso
The first thing we note is that having a natural family of invertible morphisms gives isomorphisms in the respective functor categories:
: (F F∘ F⁻¹) [D,D].≅ Id
F∘F⁻¹≅Id =
F∘F⁻¹≅Id .invertible→iso counit
[D,D](invertible→invertibleⁿ _ counit-iso)
: Id [C,C].≅ (F⁻¹ F∘ F)
Id≅F⁻¹∘F =
Id≅F⁻¹∘F .invertible→iso unit
[C,C](invertible→invertibleⁿ _ unit-iso)
= [C,C]._≅_.from Id≅F⁻¹∘F
unit⁻¹ = [D,D]._≅_.from F∘F⁻¹≅Id counit⁻¹
We chose, for definiteness, the above definition of equivalence of categories, since it provides convenient access to the most useful data: The induced natural isomorphisms, the adjunction unit/counit, and the triangle identities. It is a lot of data to come up with by hand, though, so we provide some alternatives:
Fully faithful, essentially surjective🔗
Any fully faithful and (split!) essentially surjective functor determines an equivalence of precategories. Recall that “split essentially surjective” means we have some determined procedure for picking out an essential fibre over any object : an object together with a specified isomorphism .
module _ {F : Functor C D} (ff : is-fully-faithful F) (eso : is-split-eso F) where
import Cat.Reasoning C as C
import Cat.Reasoning D as D
private module di = D._≅_
private
: ∀ {x y} → D.Hom (F .F₀ x) (F .F₀ y) → C.Hom _ _
ff⁻¹ = equiv→inverse ff
ff⁻¹ module ff {x} {y} = Equiv (_ , ff {x} {y})
It remains to show that, when is fully faithful, this assignment of essential fibres extends to a functor . For the object part, we send to the specified preimage. For the morphism part, the splitting gives us isomorphisms and , so that we may form the composite ; Fullness then completes the construction.
: Functor D C
ff+split-eso→inverse .F₀ x = eso x .fst
ff+split-eso→inverse .F₁ {x} {y} f =
ff+split-eso→inverse (f*y-iso .D._≅_.from D.∘ f D.∘ f*x-iso .D._≅_.to)
ff⁻¹ where
open Σ (eso x) renaming (fst to f*x ; snd to f*x-iso)
open Σ (eso y) renaming (fst to f*y ; snd to f*y-iso)
We must then, as usual, prove that this definition preserves identities and distributes over composites, so that we really have a functor. Preservation of identities is immediate; Distribution over composites is by faithfulness.
.F-id {x} =
ff+split-eso→inverse (f*x-iso .di.from D.∘ ⌜ D.id D.∘ f*x-iso .di.to ⌝) ≡⟨ ap! (D.idl _) ⟩
ff⁻¹ (f*x-iso .di.from D.∘ f*x-iso .di.to) ≡⟨ ap ff⁻¹ (f*x-iso .di.invr) ⟩
ff⁻¹ .id ≡˘⟨ ap ff⁻¹ (F-id F) ⟩
ff⁻¹ D(F₁ F C.id) ≡⟨ ff.η _ ⟩
ff⁻¹ .id ∎
Cwhere open Σ (eso x) renaming (fst to f*x ; snd to f*x-iso)
.F-∘ {x} {y} {z} f g =
ff+split-eso→inverse {F = F} ff (
fully-faithful→faithful (ff⁻¹ (ffz D.∘ (f D.∘ g) D.∘ ftx)) ≡⟨ ff.ε _ ⟩
F₁ F .∘ (f D.∘ g) D.∘ ftx ≡⟨ cat! D ⟩
ffz D.∘ f D.∘ ⌜ D.id ⌝ D.∘ g D.∘ ftx ≡˘⟨ ap¡ (f*y-iso .di.invl) ⟩
ffz D.∘ f D.∘ (fty D.∘ ffy) D.∘ g D.∘ ftx ≡⟨ cat! D ⟩
ffz D(ffz D.∘ f D.∘ fty) D.∘ (ffy D.∘ g D.∘ ftx) ≡˘⟨ ap₂ D._∘_ (ff.ε _) (ff.ε _) ⟩
(ff⁻¹ _) D.∘ F₁ F (ff⁻¹ _) ≡˘⟨ F-∘ F _ _ ⟩
F₁ F (ff⁻¹ _ C.∘ ff⁻¹ _) ∎
F₁ F )
where
open Σ (eso x) renaming (fst to f*x ; snd to f*x-iso)
open Σ (eso y) renaming (fst to f*y ; snd to f*y-iso)
open Σ (eso z) renaming (fst to f*z ; snd to f*z-iso)
= f*z-iso .di.from
ffz = f*z-iso .di.to
ftz = f*y-iso .di.from
ffy = f*y-iso .di.to
fty = f*x-iso .di.from
ffx = f*x-iso .di.to ftx
We will, for brevity, refer to the functor we’ve just built as
,
rather than its “proper name” ff+split-eso→inverse
. Hercules now only
has 11 labours to go: We must construct unit and counit natural
transformations, prove that they satisfy the triangle identities, and
prove that the unit/counit we define are componentwise invertible. I’ll
keep the proofs of naturality in <details>
tags
since.. they’re rough.
private
= ff+split-eso→inverse G
For the unit, we have an object and we’re asked to provide a morphism — where, recall, the notation represents the chosen essential fibre of over . By fullness, it suffices to provide a morphism ; But recall that the essential fibre comes equipped with an isomorphism .
: Id => (G F∘ F)
ff+split-eso→unit .η x = ff⁻¹ (f*x-iso .di.from)
ff+split-eso→unit where open Σ (eso (F₀ F x)) renaming (fst to f*x ; snd to f*x-iso)
Naturality of ff+split-eso→unit
.
.is-natural x y f =
ff+split-eso→unit {F = F} ff (
fully-faithful→faithful (ff⁻¹ ffy C.∘ f) ≡⟨ F-∘ F _ _ ⟩
F₁ F (ff⁻¹ ffy) ⌝ D.∘ F₁ F f ≡⟨ ap! (ff.ε _) ⟩
⌜ F₁ F .∘ ⌜ F₁ F f ⌝ ≡⟨ ap! (sym (D.idr _) ∙ ap (F₁ F f D.∘_) (sym (f*x-iso .di.invl))) ⟩
ffy D.∘ F₁ F f D.∘ ftx D.∘ ffx ≡⟨ cat! D ⟩
ffy D(ffy D.∘ F₁ F f D.∘ ftx) D.∘ ffx ≡˘⟨ ap₂ D._∘_ (ff.ε _) (ff.ε _) ⟩
(ff⁻¹ (ffy D.∘ F₁ F f D.∘ ftx)) D.∘ F₁ F (ff⁻¹ ffx) ≡˘⟨ F-∘ F _ _ ⟩
F₁ F (ff⁻¹ (ffy D.∘ F₁ F f D.∘ ftx) C.∘ ff⁻¹ ffx) ≡⟨⟩
F₁ F (F₁ (G F∘ F) f C.∘ x→f*x) ∎
F₁ F )
where
open Σ (eso (F₀ F x)) renaming (fst to f*x ; snd to f*x-iso)
open Σ (eso (F₀ F y)) renaming (fst to f*y ; snd to f*y-iso)
= f*y-iso .di.from
ffy = f*y-iso .di.to
fty = f*x-iso .di.from
ffx = f*x-iso .di.to
ftx
: C.Hom x f*x
x→f*x = ff⁻¹ (f*x-iso .di.from)
x→f*x
: C.Hom y f*y
y→f*y = ff⁻¹ (f*y-iso .di.from) y→f*y
For the counit, we have to provide a morphism ; We can again pick the given isomorphism.
: (F F∘ G) => Id
ff+split-eso→counit .η x = f*x-iso .di.to
ff+split-eso→counit where open Σ (eso x) renaming (fst to f*x ; snd to f*x-iso)
Naturality of ff+split-eso→counit
.is-natural x y f =
ff+split-eso→counit .∘ ⌜ F₁ F (ff⁻¹ (ffy D.∘ f D.∘ ftx)) ⌝ ≡⟨ ap! (ff.ε _) ⟩
fty D.∘ ffy D.∘ f D.∘ ftx ≡⟨ D.cancell (f*y-iso .di.invl) ⟩
fty D.∘ ftx ∎
f Dwhere
open Σ (eso x) renaming (fst to f*x ; snd to f*x-iso)
open Σ (eso y) renaming (fst to f*y ; snd to f*y-iso)
= f*y-iso .di.from
ffy = f*y-iso .di.to
fty = f*x-iso .di.to ftx
Checking the triangle identities, and that the adjunction unit/counit defined above are natural isomorphisms, is routine. We present the calculations without commentary:
open _⊣_
: F ⊣ G
ff+split-eso→F⊣inverse .unit = ff+split-eso→unit
ff+split-eso→F⊣inverse .counit = ff+split-eso→counit
ff+split-eso→F⊣inverse .zig {x} =
ff+split-eso→F⊣inverse .∘ F₁ F (ff⁻¹ ffx) ≡⟨ ap (ftx D.∘_) (ff.ε _) ⟩
ftx D.∘ ffx ≡⟨ f*x-iso .di.invl ⟩
ftx D.id ∎ D
The zag
identity needs an appeal
to faithfulness:
.zag {x} =
ff+split-eso→F⊣inverse {F = F} ff (
fully-faithful→faithful (ff⁻¹ (ffx D.∘ ftx D.∘ fftx) C.∘ ff⁻¹ fffx) ≡⟨ F-∘ F _ _ ⟩
F₁ F (ff⁻¹ (ffx D.∘ ftx D.∘ fftx)) D.∘ F₁ F (ff⁻¹ fffx) ≡⟨ ap₂ D._∘_ (ff.ε _) (ff.ε _) ⟩
F₁ F (ffx D.∘ ftx D.∘ fftx) D.∘ fffx ≡⟨ cat! D ⟩
(ffx D.∘ ftx) D.∘ (fftx D.∘ fffx) ≡⟨ ap₂ D._∘_ (f*x-iso .di.invr) (f*f*x-iso .di.invl) ⟩
.id D.∘ D.id ≡⟨ D.idl _ ∙ sym (F-id F) ⟩
D.id ∎
F₁ F C)
Now to show they are componentwise invertible:
open is-equivalence
: is-equivalence F
ff+split-eso→is-equivalence .F⁻¹ = G
ff+split-eso→is-equivalence .F⊣F⁻¹ = ff+split-eso→F⊣inverse
ff+split-eso→is-equivalence .counit-iso x = record
ff+split-eso→is-equivalence { inv = f*x-iso .di.from
; inverses = record
{ invl = f*x-iso .di.invl
; invr = f*x-iso .di.invr }
}
where open Σ (eso x) renaming (fst to f*x ; snd to f*x-iso)
Since the unit is defined in terms of fullness, showing it is invertible needs an appeal to faithfulness (two, actually):
.unit-iso x = record
ff+split-eso→is-equivalence { inv = ff⁻¹ (f*x-iso .di.to)
; inverses = record
{ invl = fully-faithful→faithful {F = F} ff (
(ff⁻¹ ffx C.∘ ff⁻¹ ftx) ≡⟨ F-∘ F _ _ ⟩
F₁ F (ff⁻¹ ffx) D.∘ F₁ F (ff⁻¹ ftx) ≡⟨ ap₂ D._∘_ (ff.ε _) (ff.ε _) ⟩
F₁ F .∘ ftx ≡⟨ f*x-iso .di.invr ⟩
ffx D.id ≡˘⟨ F-id F ⟩
D.id ∎)
F₁ F C; invr = fully-faithful→faithful {F = F} ff (
(ff⁻¹ ftx C.∘ ff⁻¹ ffx) ≡⟨ F-∘ F _ _ ⟩
F₁ F (ff⁻¹ ftx) D.∘ F₁ F (ff⁻¹ ffx) ≡⟨ ap₂ D._∘_ (ff.ε _) (ff.ε _) ⟩
F₁ F .∘ ffx ≡⟨ f*x-iso .di.invl ⟩
ftx D.id ≡˘⟨ F-id F ⟩
D.id ∎)
F₁ F C}
}
where
open Σ (eso (F₀ F x)) renaming (fst to f*x ; snd to f*x-iso)
= f*x-iso .di.from
ffx = f*x-iso .di.to ftx
Between categories🔗
Above, we made an equivalence out of any fully faithful and split essentially surjective functor. In set-theoretic mathematics (and indeed between strict categories), the splitting condition can not be lifted constructively: the statement “every (ff, eso) functor between strict categories is an equivalence” is equivalent to the axiom of choice.
However, between univalent categories, the situation is different: Any essentially surjective fully faithful functor splits. In particular, any ff functor between univalent categories has propositional essential fibres, so a “mere” essential surjection is automatically split. However, note that both the domain and codomain have to be categories for the argument to go through.
module
_ (F : Functor C D) (ccat : is-category C) (dcat : is-category D)
(ff : is-fully-faithful F)
where
private
module C = Cat.Reasoning C
module D = Cat.Reasoning D
So, suppose we have categories and , together with a fully faithful functor . For any , we’re given an inhabitant of , which we want to “get out” from under the truncation. For this, we’ll show that the type being truncated is a proposition, so that we may “untruncate” it.
: ∀ y → is-prop (Essential-fibre F y)
Essential-fibre-between-cats-is-prop (x , i) (y , j) = they're-equal where Essential-fibre-between-cats-is-prop z
For this magic trick, assume we’re given a
,
together with objects
and isomorphisms
and
.
We must show that
,
and that over this path,
.
Since
is fully faithful, we can find an isomorphism
in
,
which
sends back to
.
: F₀ F x D.≅ F₀ F y
Fx≅Fy = i D.∘Iso (j D.Iso⁻¹)
Fx≅Fy
: x C.≅ y
x≅y = is-ff→essentially-injective {F = F} ff Fx≅Fy x≅y
Furthermore, since we’re working with categories, these isomorphisms restrict to paths and . We’re half-done: we’ve shown that some exists, and it remains to show that over this path we have . More specifically, we must give a path laying over .
: x ≡ y
x≡y = ccat .to-path x≅y
x≡y
: F₀ F x ≡ F₀ F y
Fx≡Fy = dcat .to-path Fx≅Fy Fx≡Fy
Rather than showing it over
directly, we’ll show it over the path
we constructed independently. This is because we can use the helper
Hom-pathp-reflr-iso
to establish
the result with far less computation:
: PathP (λ i → Fx≡Fy i D.≅ z) i j
over' = D.≅-pathp Fx≡Fy refl
over' (Univalent.Hom-pathp-refll-iso dcat (D.cancell (i .D._≅_.invl)))
We must then connect with this path . But since we originally got by full faithfulness of , they are indeed the same path:
abstract
: ap (F₀ F) x≡y ≡ Fx≡Fy
square =
square (F₀ F) x≡y ≡⟨ F-map-path ccat dcat F x≅y ⟩
ap .to-path ⌜ F-map-iso F x≅y ⌝ ≡⟨ ap! (equiv→counit (is-ff→F-map-iso-is-equiv {F = F} ff) _) ⟩
dcat .to-path Fx≅Fy ∎
dcat
: PathP (λ i → F₀ F (x≡y i) D.≅ z) i j
over = transport (λ l → PathP (λ m → square (~ l) m D.≅ z) i j) over' over
Hence — blink and you’ll miss it — the essential fibres of over any are propositions, so it suffices for them to be merely inhabited for the functor to be split eso. With tongue firmly in cheek we call this result the theorem of choice.
= Σ-pathp x≡y over
they're-equal
: is-eso F → is-split-eso F
Theorem-of-choice =
Theorem-of-choice eso y (λ _ → Essential-fibre-between-cats-is-prop y)
∥-∥-elim (λ x → x) (eso y)
This theorem implies that any fully faithful, “merely” essentially surjective functor between categories is an equivalence:
: is-eso F → is-equivalence F
ff+eso→is-equivalence = ff+split-eso→is-equivalence ff (Theorem-of-choice eso) ff+eso→is-equivalence eso
Furthermore, if is an equivalence between categories, then it’s an equivalence-on-objects functor. The inverse functor gives us a way to turn objects of back into objects of , and unit/counit of the equivalence ensure that and , so all that remains is to use the fact that and are categories to get the requisite paths.
is-cat-equivalence→equiv-on-objects: ∀ {F : Functor C D}
→ (ccat : is-category C) (dcat : is-category D)
→ is-equivalence F → is-equiv-on-objects F
{C = C} {D = D} {F = F} ccat dcat eqv =
is-cat-equivalence→equiv-on-objects
is-iso→is-equiv $(e.F⁻¹ .F₀)
iso (λ d → dcat .to-path (D.invertible→iso _ (e.counit-iso d)))
(λ c → sym $ ccat .to-path (C.invertible→iso _ (e.unit-iso c)))
where
module C = Cat.Reasoning C
module D = Cat.Reasoning D
module e = is-equivalence eqv
Isomorphisms🔗
Another, more direct way of proving that a functor is an equivalence of precategories is proving that it is an isomorphism of precategories: It’s fully faithful, thus a typal equivalence of morphisms, and in addition its action on objects is an equivalence of types.
record is-precat-iso (F : Functor C D) : Type (adj-level C D) where
no-eta-equality
constructor iso
field
: is-fully-faithful F
has-is-ff : is-equiv (F₀ F) has-is-iso
Such a functor is (immediately) fully faithful, and the inverse has-is-iso
means that it is split
essentially surjective; For given
,
the inverse of
gives us an object
;
We must then provide an isomorphism
,
but those are identical, hence isomorphic.
module _ {F : Functor C D} (p : is-precat-iso F) where
open is-precat-iso p
: is-split-eso F
is-precat-iso→is-split-eso = equiv→inverse has-is-iso ob , isom
is-precat-iso→is-split-eso ob where isom = path→iso {C = D} (equiv→counit has-is-iso _)
Thus, by the theorem above, is an adjoint equivalence of precategories.
: is-equivalence F
is-precat-iso→is-equivalence =
is-precat-iso→is-equivalence ff+split-eso→is-equivalence has-is-ff is-precat-iso→is-split-eso
Properties of equivalences🔗
If is fully-faithfuly and essentially surjective, then for every hom-set there (merely) exists an equivalent hom-set .
ff+split-eso→hom-equiv: (F : Functor C D)
→ is-fully-faithful F
→ is-split-eso F
→ ∀ (d d' : D.Ob) → Σ[ c ∈ C.Ob ] Σ[ c' ∈ C.Ob ] (C.Hom c c' ≃ D.Hom d d')
=
ff+split-eso→hom-equiv F ff split-eso d d' .fst , d'-fib .fst ,
d-fib (F .F₁ , ff) ∙e D.iso→hom-equiv (d-fib .snd) (d'-fib .snd)
where
= split-eso d
d-fib = split-eso d'
d'-fib
ff+eso→hom-equiv: (F : Functor C D)
→ is-fully-faithful F
→ is-eso F
→ ∀ (d d' : D.Ob) → ∥ Σ[ c ∈ C.Ob ] Σ[ c' ∈ C.Ob ] (C.Hom c c' ≃ D.Hom d d') ∥
= do
ff+eso→hom-equiv F ff eso d d' (c , Fc≅d) ← eso d
(c' , Fc'≅d') ← eso d'
(c , c' , (F .F₁ , ff) ∙e D.iso→hom-equiv Fc≅d Fc'≅d') pure
This allows us to prove a very useful little lemma: if is a fully-faithful, essentially surjective functor, then any property of hom-sets of that holds for all hom-sets must also hold for all hom-sets of .
ff+eso→preserves-hom-props: ∀ {ℓ} (F : Functor C D)
→ is-fully-faithful F
→ is-eso F
→ (P : Type (ℓc ⊔ ℓd) → Type ℓ)
→ (∀ A → is-prop (P A))
→ (∀ c c' → P (Lift ℓd (C.Hom c c')))
→ ∀ d d' → P (Lift ℓc (D.Hom d d'))
=
ff+eso→preserves-hom-props F ff eso P prop P-hom d d' (prop (Lift ℓc (D.Hom d d'))) $ do
∥-∥-proj (c , c' , eqv) ← ff+eso→hom-equiv F ff eso d d'
(transport (ap P (ua (Lift-ap eqv))) (P-hom c c')) pure
As a corollary, we note that if is a fully-faithful, essentially surjective functor, then if the hom-sets of are all -types, then the hom-sets of must also be -types.
ff+eso→hom-hlevel: ∀ {n} (F : Functor C D)
→ is-fully-faithful F
→ is-eso F
→ (∀ c c' → is-hlevel (C.Hom c c') n)
→ ∀ d d' → is-hlevel (D.Hom d d') n
{n = n} F ff eso C-hlevel d d' =
ff+eso→hom-hlevel _ $
Lift-is-hlevel'
ff+eso→preserves-hom-props F ff eso(λ A → is-hlevel A n) (λ _ → is-hlevel-is-prop n)
(λ c c' → Lift-is-hlevel n (C-hlevel c c')) d d'
Note that if is fully faithful and split essentially surjective, then we can drop the requirement that must be a prop.
ff+split-eso→preserves-hom-fams: ∀ {ℓ} (F : Functor C D)
→ is-fully-faithful F
→ is-split-eso F
→ (P : Type (ℓc ⊔ ℓd) → Type ℓ)
→ (∀ c c' → P (Lift ℓd (C.Hom c c')))
→ ∀ d d' → P (Lift ℓc (D.Hom d d'))
=
ff+split-eso→preserves-hom-fams F ff split-eso P P-hom d d'
transport(ap P (ua (Lift-ap (ff+split-eso→hom-equiv F ff split-eso d d' .snd .snd))))
(P-hom _ _)
As a corollary, equivalences preserve all families over hom sets.
equivalence→preserves-hom-fams: ∀ {ℓ} (E : Equivalence C D)
→ (P : Type (ℓc ⊔ ℓd) → Type ℓ)
→ (∀ c c' → P (Lift ℓd (C.Hom c c')))
→ ∀ d d' → P (Lift ℓc (D.Hom d d'))
=
equivalence→preserves-hom-fams E (Equivalence.To E)
ff+split-eso→preserves-hom-fams (is-equivalence→is-ff _ (Equivalence.To-equiv E))
(is-equivalence→is-split-eso _ (Equivalence.To-equiv E))
Being an equivalence is also invariant under natural isomorphism.
is-equivalence-natural-iso: ∀ {F G : Functor C D}
→ F ≅ⁿ G
→ is-equivalence F → is-equivalence G
{C = C} {D = D} {F = F} {G = G} α F-eqv = G-eqv where
is-equivalence-natural-iso open is-equivalence
module C = Cat.Reasoning C
module D = Cat.Reasoning D
: is-equivalence G
G-eqv .F⁻¹ = F-eqv .F⁻¹
G-eqv .F⊣F⁻¹ = adjoint-natural-isol α (F-eqv .F⊣F⁻¹)
G-eqv .unit-iso x =
G-eqv .invertible-∘
C(C.invertible-∘
(F-map-invertible (F-eqv .F⁻¹) (isoⁿ→is-invertible α x))
.id-invertible)
C(F-eqv .unit-iso x)
.counit-iso x =
G-eqv .invertible-∘
D(F-eqv .counit-iso x)
(D.invertible-∘
(F-map-invertible F C.id-invertible)
(isoⁿ→is-invertible α _ D.invertible⁻¹))
Equivalences are invertible.
_Equivalence⁻¹
: Equivalence C D → Equivalence D C
(E Equivalence⁻¹) .Equivalence.To = Equivalence.From E
(E Equivalence⁻¹) .Equivalence.To-equiv = Equivalence.inverse-equivalence E
Equivalences are also composable, as adjoints compose.
is-equivalence-∘: ∀ {F : Functor D E} {G : Functor C D}
→ is-equivalence F → is-equivalence G
→ is-equivalence (F F∘ G)
{E = E} {C = C} {F = F} {G = G} F-eqv G-eqv = FG-eqv where
is-equivalence-∘ module F-eqv = is-equivalence F-eqv
module G-eqv = is-equivalence G-eqv
module C = Cat.Reasoning C
module E = Cat.Reasoning E
: is-equivalence (F F∘ G)
FG-eqv .F⁻¹ = G-eqv.F⁻¹ F∘ F-eqv.F⁻¹
FG-eqv .F⊣F⁻¹ = LF⊣GR G-eqv.F⊣F⁻¹ F-eqv.F⊣F⁻¹
FG-eqv .unit-iso x =
FG-eqv .invertible-∘
C(F-map-invertible G-eqv.F⁻¹ (F-eqv.unit-iso (G .F₀ x)))
(G-eqv.unit-iso x)
.counit-iso x =
FG-eqv .invertible-∘
E(F-eqv.counit-iso x)
(F-map-invertible F (G-eqv.counit-iso (F-eqv .F⁻¹ .F₀ x)))
_∘Equivalence_ : Equivalence C D → Equivalence D E → Equivalence C E
(F ∘Equivalence G) .Equivalence.To =
.To G F∘ Equivalence.To F
Equivalence(F ∘Equivalence G) .Equivalence.To-equiv =
(Equivalence.To-equiv G) (Equivalence.To-equiv F) is-equivalence-∘