module Cat.Functor.Properties whereFunctorsđ
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:
is-full : Functor C D â Type _
is-full {C = C} {D = D} F =
â {x y} (g : D .Hom (Fâ F x) (Fâ F y)) â â[ f â C .Hom x y ] (Fâ F f ⥠g)A functor is faithful when its action on hom-sets is injective:
is-faithful : Functor C D â Type _
is-faithful F = â {x y} â injective (Fâ F {x = x} {y})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.
is-fully-faithful : Functor C D â Type _
is-fully-faithful F = â {x y} â is-equiv (Fâ F {x = x} {y})
fully-faithfulâfaithful : {F : Functor C D} â is-fully-faithful F â is-faithful F
fully-faithfulâfaithful f = Equiv.injective (_ , f)
fully-faithfulâfull : {F : Functor C D} â is-fully-faithful F â is-full F
fully-faithfulâfull {F = F} ff g = inc (equivâinverse ff g , equivâcounit ff g)
full+faithfulâff
: (F : Functor C D) â is-full F â is-faithful F â is-fully-faithful F
full+faithfulâff {C = C} {D = D} F surj inj .is-eqv = p where
img-is-prop : â {x y} f â is-prop (fibre (Fâ F {x = x} {y}) f)
img-is-prop f (g , p) (h , q) = ÎŁ-prop-path (λ _ â D .Hom-set _ _ _ _) (inj (p â sym q))
p : â {x y} f â is-contr (fibre (Fâ F {x = x} {y}) f)
p f .centre = â„-â„-elim (λ _ â img-is-prop f) (λ x â x) (surj f)
p f .paths = img-is-prop 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
is-ffâis-conservative {F = F} ff f isinv = i where
open Cm.is-invertible
open Cm.InversesSince the functor is ff, we can find a map ââ in the domain category to serve as an inverse for :
g : C.Hom _ _
g = equivâinverse ff (isinv .Dm.is-invertible.inv)
module ff {a} {b} = Equiv (_ , ff {a} {b})
Ffog =
Fâ F (f C.â g) âĄâš F-â F _ _ â©
Fâ F f D.â Fâ F g âĄâš apâ D._â_ refl (ff.Δ _) â isinv .Dm.is-invertible.invl â©
D.id â
Fgof =
Fâ F (g C.â f) âĄâš F-â F _ _ â©
Fâ F g D.â Fâ F f âĄâš apâ D._â_ (ff.Δ _) refl â isinv .Dm.is-invertible.invr â©
D.id â
i : Cm.is-invertible _
i .inv = g
i .inverses .invl =
f C.â g âĄâš sym (ff.η _) â©
ff.from â Fâ F (f C.â g) â âĄâš ap! (Ffog â sym (F-id F)) â©
ff.from (Fâ F C.id) âĄâš ff.η _ â©
C.id â
i .inverses .invr =
g C.â f âĄâš sym (ff.η _) â©
ff.from â Fâ F (g C.â f) â âĄâš ap! (Fgof â sym (F-id F)) â©
ff.from (Fâ F C.id) âĄâš ff.η _ â©
C.id â
is-ffâessentially-injective
: {F : Functor C D} â is-fully-faithful F
â â {X Y} â Fâ F X Dm.â
Fâ F Y
â X Cm.â
Y
is-ffâessentially-injective {F = F} ff im = im' where
-- Cm.make-iso (equivâinverse ff to) inv invl invr
open Dm._â
_ im using (to ; from ; inverses)
D-inv' : Dm.is-invertible (Fâ F (equivâinverse ff to))
D-inv' .Dm.is-invertible.inv = from
D-inv' .Dm.is-invertible.inverses =
subst (λ e â Dm.Inverses e from) (sym (equivâcounit ff _)) inverses
open Cm.is-invertible (is-ffâis-conservative {F = F} ff (equivâinverse ff to) D-inv')
im' : _ Cm.â
_
im' .to = equivâinverse ff to
im' .from = inv
im' .inverses .Cm.Inverses.invl = invl
im' .inverses .Cm.Inverses.invr = invrEssential fibresđ
The essential fibre of a functor over an object is the space of objects of which takes, up to isomorphism, to .
Essential-fibre : Functor C D â D .Ob â Type _
Essential-fibre {D = D} F y = ÎŁ[ x â _ ] (Fâ F x â
y)
where open import Cat.Morphism DA 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:
is-split-eso : Functor C D â Type _
is-split-eso F = â y â Essential-fibre F y
is-eso : Functor C D â Type _
is-eso F = â y â â„ Essential-fibre F y â„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.
is-full-on-isos : Functor C D â Type (o â h â hâ)
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
faithful : is-faithful F
isos-full : is-full-on-isos F
open is-pseudomonicSomewhat 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
pseudomonicâconservative {F = F} pseudo {x} {y} f inv =
â„-â„-rec C.is-invertible-is-prop
(λ (g , p) â
C.make-invertible (C.from g)
(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
pseudomonicâessentially-injective {F = F} pseudo f =
â„-â„-rec (faithfulâiso-fibre-prop F (pseudo .faithful) f)
(λ x â x)
(pseudo .isos-full f) .fstFully faithful functors are pseudomonic, as they are faithful and essentially injective.
ffâpseudomonic
: â {F : Functor C D}
â is-fully-faithful F
â is-pseudomonic F
ffâpseudomonic {F} ff .faithful = fully-faithfulâfaithful {F = F} ff
ffâpseudomonic {F} ff .isos-full f =
inc (is-ffâessentially-injective {F = F} ff f ,
ext (equivâcounit ff (D.to f)))Equivalence on objects functorsđ
A functor is an equivalence on objects if its action on objects is an equivalence.
is-equiv-on-objects : (F : Functor C D) â Type _
is-equiv-on-objects F = is-equiv (F .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
equiv-on-objectsâsplit-eso {D = D} F eqv y =
equivâinverse eqv y , pathâiso (equivâcounit eqv y)
equiv-on-objectsâeso : â (F : Functor C D) â is-equiv-on-objects F â is-eso F
equiv-on-objectsâeso F eqv y = inc (equiv-on-objectsâsplit-eso F eqv y)