module Cat.Functor.Properties where
private variable
: Level
o h o₁ h₁ : Precategory o h
B C D open Precategory
open Functor
Functors🔗
This module defines the most important clases of functors: Full, faithful, fully faithful (abbreviated ff), split essentially surjective and (“merely”) essentially surjective.
A functor is full when its action on hom-sets is surjective:
: Functor C D → Type _
is-full {C = C} {D = D} F = ∀ {x y} → is-surjective (F .F₁ {x = x} {y}) is-full
A functor is faithful when its action on hom-sets is injective:
: Functor C D → Type _
is-faithful = ∀ {x y} → injective (F .F₁ {x = x} {y}) is-faithful F
module _ {C : Precategory o h} {D : Precategory o₁ h₁} where
private module _ where
module C = Cat.Reasoning C
module D = Cat.Reasoning D
open Cat.Reasoning using (_≅_ ; Inverses)
open _≅_ public
open Inverses public
faithful→iso-fibre-prop: ∀ (F : Functor C D)
→ is-faithful F
→ ∀ {x y} → (f : F # x D.≅ F # y)
→ is-prop (Σ[ g ∈ x C.≅ y ] (F-map-iso F g ≡ f))
(g , p) (g' , q) =
faithful→iso-fibre-prop F faithful f (faithful (ap D.to (p ∙ sym q)))
Σ-prop-path! $ ext
is-faithful-∘: ∀ {F : Functor C D} {G : Functor B C}
→ is-faithful F → is-faithful G
→ is-faithful (F F∘ G)
= Gf (Ff p) is-faithful-∘ Ff Gf p
Fully faithful functors🔗
A functor is fully faithful (abbreviated ff) when its action on hom-sets is an equivalence. Since Hom-sets are sets, it suffices for the functor to be full and faithful; Rather than taking this conjunction as a definition, we use the more directly useful data as a definition and prove the conjunction as a theorem.
: Functor C D → Type _
is-fully-faithful = ∀ {x y} → is-equiv (F .F₁ {x = x} {y})
is-fully-faithful F
: {F : Functor C D} → is-fully-faithful F → is-faithful F
ff→faithful = Equiv.injective (_ , f)
ff→faithful f
: {F : Functor C D} → is-fully-faithful F → is-full F
ff→full {F = F} ff g = inc (equiv→inverse ff g , equiv→counit ff g)
ff→full
full+faithful→ff: (F : Functor C D) → is-full F → is-faithful F → is-fully-faithful F
{C = C} {D = D} F surj inj .is-eqv = p where
full+faithful→ff : ∀ {x y} f → is-prop (fibre (F .F₁ {x = x} {y}) f)
img-is-prop (g , p) (h , q) = Σ-prop-path (λ _ → D .Hom-set _ _ _ _) (inj (p ∙ sym q))
img-is-prop f
: ∀ {x y} f → is-contr (fibre (F .F₁ {x = x} {y}) f)
p .centre = ∥-∥-elim (λ _ → img-is-prop f) (λ x → x) (surj f)
p f .paths = img-is-prop f _ p f
A very important property of fully faithful functors (like is that they are conservative: If the image of under is an isomorphism then was really an isomorphism
module _ {C : Precategory o h} {D : Precategory o₁ h₁} where
private
module C = Precategory C
module D = Precategory D
import Cat.Morphism C as Cm
import Cat.Morphism D as Dm
is-ff→is-conservative: {F : Functor C D} → is-fully-faithful F
→ ∀ {X Y} (f : C.Hom X Y) → Dm.is-invertible (F .F₁ f)
→ Cm.is-invertible f
{F = F} ff f isinv = i where
is-ff→is-conservative open Cm.is-invertible
open Cm.Inverses
Since the functor is ff, we can find a map “” in the domain category to serve as an inverse for
: C.Hom _ _
g = equiv→inverse ff (isinv .Dm.is-invertible.inv)
g module ff {a} {b} = Equiv (_ , ff {a} {b})
=
Ffog .F₁ (f C.∘ g) ≡⟨ F .F-∘ _ _ ⟩
F .F₁ f D.∘ F .F₁ g ≡⟨ ap₂ D._∘_ refl (ff.ε _) ∙ isinv .Dm.is-invertible.invl ⟩
F .id ∎
D
=
Fgof .F₁ (g C.∘ f) ≡⟨ F .F-∘ _ _ ⟩
F .F₁ g D.∘ F .F₁ f ≡⟨ ap₂ D._∘_ (ff.ε _) refl ∙ isinv .Dm.is-invertible.invr ⟩
F .id ∎
D
: Cm.is-invertible _
i .inv = g
i .inverses .invl =
i .∘ g ≡⟨ sym (ff.η _) ⟩
f C.from ⌜ F .F₁ (f C.∘ g) ⌝ ≡⟨ ap! (Ffog ∙ sym (F .F-id)) ⟩
ff.from (F .F₁ C.id) ≡⟨ ff.η _ ⟩
ff.id ∎
C.inverses .invr =
i .∘ f ≡⟨ sym (ff.η _) ⟩
g C.from ⌜ F .F₁ (g C.∘ f) ⌝ ≡⟨ ap! (Fgof ∙ sym (F .F-id)) ⟩
ff.from (F .F₁ C.id) ≡⟨ ff.η _ ⟩
ff.id ∎
C
is-ff→essentially-injective: {F : Functor C D} → is-fully-faithful F
→ ∀ {X Y} → F .F₀ X Dm.≅ F .F₀ Y
→ X Cm.≅ Y
{F = F} ff im = im' where
is-ff→essentially-injective open Dm._≅_ im using (to ; from ; inverses)
: Dm.is-invertible (F .F₁ (equiv→inverse ff to))
D-inv' .Dm.is-invertible.inv = from
D-inv' .Dm.is-invertible.inverses =
D-inv' (λ e → Dm.Inverses e from) (sym (equiv→counit ff _)) inverses
subst
open Cm.is-invertible (is-ff→is-conservative {F = F} ff (equiv→inverse ff to) D-inv')
: _ Cm.≅ _
im' .to = equiv→inverse ff to
im' .from = inv
im' .inverses .Cm.Inverses.invl = invl
im' .inverses .Cm.Inverses.invr = invr im'
Essential fibres🔗
The essential fibre of a functor over an object is the space of objects of which takes, up to isomorphism, to
: Functor C D → D .Ob → Type _
Essential-fibre {C = C} {D = D} F y = Σ[ x ∈ C ] (F # x ≅ y)
Essential-fibre where open import Cat.Morphism D
A functor is split essentially surjective (abbreviated split eso) if there is a procedure for finding points in the essential fibre over any object. It’s essentially surjective if this procedure merely, i.e. truncatedly, finds a point:
: Functor C D → Type _
is-split-eso = ∀ y → Essential-fibre F y
is-split-eso F
: Functor C D → Type _
is-eso = ∀ y → ∥ Essential-fibre F y ∥ is-eso F
module _ {C : Precategory o h} {D : Precategory o₁ h₁} where
import Cat.Reasoning C as C
import Cat.Reasoning D as D
private module _ where
open import Cat.Reasoning using (_≅_ ; Inverses)
open _≅_ public
open Inverses public
is-ff→F-map-iso-is-equiv: {F : Functor C D} → is-fully-faithful F
→ ∀ {X Y} → is-equiv (F-map-iso F {x = X} {Y})
{F = F} ff = is-iso→is-equiv isom where
is-ff→F-map-iso-is-equiv : is-iso _
isom .is-iso.inv = is-ff→essentially-injective {F = F} ff
isom .is-iso.rinv x = ext (equiv→counit ff _)
isom .is-iso.linv x = ext (equiv→unit ff _) isom
If a functor is essentially surjective, then the precomposition functor is faithful for every precategory
is-eso→precompose-is-faithful: ∀ {oa ℓa} {A : Precategory oa ℓa}
→ (F : Functor C D)
→ is-eso F
→ is-faithful (precompose F {D = A})
This is rather abstract, so let’s unfold the statement a bit. If precomposition by is fully faithful, then we can determine equality of natural transformations between functors by only looking at the components of the form for every
Intuitively, this should be true for essentially surjective functors, as doesn’t miss out on any objects of but the actual proof involves some rather tedious calculations.
{A = A} F F-eso {G} {H} {α} {β} αL=βL =
is-eso→precompose-is-faithful λ d → ∥-∥-out! do
ext (c , Fc≅d) ← F-eso d
let module Fc≅d = D._≅_ Fc≅d
pure $.η d ≡⟨ A.intror (G.annihilate (D.invl Fc≅d)) ⟩
α.η d A.∘ G.₁ Fc≅d.to A.∘ G.₁ Fc≅d.from ≡⟨ A.extendl (α.is-natural _ _ _) ⟩
α.₁ Fc≅d.to A.∘ ⌜ α.η (F.₀ c) ⌝ A.∘ G.₁ Fc≅d.from ≡⟨ ap! (unext αL=βL c) ⟩
H.₁ Fc≅d.to A.∘ β.η (F.₀ c) A.∘ G.₁ Fc≅d.from ≡⟨ A.extendl (sym (β.is-natural _ _ _)) ⟩
H.η d A.∘ G.₁ Fc≅d.to A.∘ G.₁ Fc≅d.from ≡⟨ A.elimr (G.annihilate (D.invl Fc≅d)) ⟩
β.η d ∎
βwhere
module A = Cat.Reasoning A
module F = Cat.Functor.Reasoning F
module G = Cat.Functor.Reasoning G
module H = Cat.Functor.Reasoning H
module α = _=>_ α
module β = _=>_ β
_ = C.epic-precomp-embedding
Another way of viewing this result is that it is a higher-dimensional
analog of the fact that precomposition with an epi is an embedding
.
Additionally, precomposition with an essentially surjective functor is conservative.
eso→precompose-conservative: ∀ {oa ℓa} {A : Precategory oa ℓa}
→ (F : Functor C D)
→ is-eso F
→ {G H : Functor D A} {α : G => H}
→ is-invertibleⁿ (α ◂ F)
→ is-invertibleⁿ α
The proof follows the same plan as the previous theorem: natural transformations are invertible when they are invertible objectwise, and covers every object of
{A = A} F F-eso {G} {H} {α} α⁻¹ =
eso→precompose-conservative λ d → ∥-∥-out! do
invertible→invertibleⁿ α (c , Fc≅d) ← F-eso d
let module Fc≅d = D._≅_ Fc≅d
pure $.make-invertible (G.₁ Fc≅d.to A.∘ α⁻¹.η c A.∘ H.₁ Fc≅d.from)
A(α.pulll (A.cancell (α⁻¹.invl #ₚ c)) ∙ H.annihilate Fc≅d.invl)
(A.pullr (α.cancelr (α⁻¹.invr #ₚ c)) ∙ G.annihilate Fc≅d.invl)
where
module A = Cat.Reasoning A
module F = Cat.Functor.Reasoning F
module G = Cat.Functor.Reasoning G
module H = Cat.Functor.Reasoning H
module α = Cat.Natural.Reasoning α
module α⁻¹ where
open is-invertibleⁿ α⁻¹ public
open Cat.Natural.Reasoning inv hiding (op) public
Pseudomonic functors🔗
A functor is pseudomonic if it is faithful and full on isomorphisms. Pseudomonic functors are arguably the correct notion of subcategory, as they ensure that we are not able to distinguish between isomorphic objects when creating a subcategory.
module _ {C : Precategory o h} {D : Precategory o₁ h₁} where
import Cat.Reasoning C as C
import Cat.Reasoning D as D
: Functor C D → Type (o ⊔ h ⊔ h₁)
is-full-on-isos =
is-full-on-isos F ∀ {x y} → (f : F .F₀ x D.≅ F .F₀ y) → ∃[ g ∈ x C.≅ y ] (F-map-iso F g ≡ f)
record is-pseudomonic (F : Functor C D) : Type (o ⊔ h ⊔ h₁) where
no-eta-equality
field
: is-faithful F
faithful : is-full-on-isos F
isos-full
open is-pseudomonic
Somewhat surprisingly, pseudomonic functors are conservative. As is full on isos, there merely exists some iso in the fibre of However, invertibility is a property of morphisms, so we can untruncate the mere existence. Once we have our hands on the isomorphism, we perform a simple calculation to note that it yields an inverse to
pseudomonic→conservative: ∀ {F : Functor C D}
→ is-pseudomonic F
→ ∀ {x y} (f : C.Hom x y) → D.is-invertible (F .F₁ f)
→ C.is-invertible f
{F = F} pseudo {x} {y} f inv =
pseudomonic→conservative .is-invertible-is-prop
∥-∥-rec C(λ (g , p) →
.make-invertible (C.from g)
C(sym (ap (C._∘ _) (pseudo .faithful (ap D.to p))) ∙ C.invl g)
(sym (ap (_ C.∘_) (pseudo .faithful (ap D.to p))) ∙ C.invr g))
(pseudo .isos-full (D.invertible→iso _ inv))
In a similar vein, pseudomonic functors are essentially injective. The proof follows a similar path to the prior one, hinging on the fact that faithful functors are an embedding on isos.
pseudomonic→essentially-injective: ∀ {F : Functor C D}
→ is-pseudomonic F
→ ∀ {x y} → F .F₀ x D.≅ F .F₀ y
→ x C.≅ y
{F = F} pseudo f =
pseudomonic→essentially-injective (faithful→iso-fibre-prop F (pseudo .faithful) f)
∥-∥-rec (λ x → x)
(pseudo .isos-full f) .fst
Fully faithful functors are pseudomonic, as they are faithful and essentially injective.
ff→pseudomonic: ∀ {F : Functor C D}
→ is-fully-faithful F
→ is-pseudomonic F
{F} ff .faithful = ff→faithful {F = F} ff
ff→pseudomonic {F} ff .isos-full f =
ff→pseudomonic (is-ff→essentially-injective {F = F} ff f ,
inc (equiv→counit ff (D.to f))) ext
Equivalence on objects functors🔗
A functor is an equivalence on objects if its action on objects is an equivalence.
: (F : Functor C D) → Type _
is-equiv-on-objects = is-equiv (F .F₀) is-equiv-on-objects F
If is an equivalence-on-objects functor, then it is (split) essentially surjective.
equiv-on-objects→split-eso: ∀ (F : Functor C D) → is-equiv-on-objects F → is-split-eso F
{D = D} F eqv y =
equiv-on-objects→split-eso (equiv→counit eqv y)
equiv→inverse eqv y , path→iso
: ∀ (F : Functor C D) → is-equiv-on-objects F → is-eso F
equiv-on-objects→eso = inc (equiv-on-objects→split-eso F eqv y) equiv-on-objects→eso F eqv y