module Cat.Morphism {o h} (C : Precategory o h) where
Morphisms🔗
This module defines the three most important classes of morphisms that can be found in a category: monomorphisms, epimorphisms, and isomorphisms.
Monos🔗
A morphism is said to be monic when it is left-cancellable. A monomorphism from to , written , is a monic morphism.
: Hom a b → Type _
is-monic {a = a} f = ∀ {c} → (g h : Hom c a) → f ∘ g ≡ f ∘ h → g ≡ h
is-monic
: ∀ {a b} (f : Hom a b) → is-prop (is-monic f)
is-monic-is-prop {c} g h p = Hom-set _ _ _ _ (x g h p) (y g h p) i
is-monic-is-prop f x y i
record _↪_ (a b : Ob) : Type (o ⊔ h) where
field
: Hom a b
mor : is-monic mor
monic
open _↪_ public
Conversely, a morphism is said to be epic when it is right-cancellable. An epimorphism from to , written , is an epic morphism.
Epis🔗
: Hom a b → Type _
is-epic {b = b} f = ∀ {c} → (g h : Hom b c) → g ∘ f ≡ h ∘ f → g ≡ h
is-epic
: ∀ {a b} (f : Hom a b) → is-prop (is-epic f)
is-epic-is-prop {c} g h p = Hom-set _ _ _ _ (x g h p) (y g h p) i
is-epic-is-prop f x y i
record _↠_ (a b : Ob) : Type (o ⊔ h) where
field
: Hom a b
mor : is-epic mor
epic
open _↠_ public
The identity morphism is monic and epic.
: ∀ {a} → is-monic (id {a})
id-monic = sym (idl _) ·· p ·· idl _
id-monic g h p
: ∀ {a} → is-epic (id {a})
id-epic = sym (idr _) ·· p ·· idr _ id-epic g h p
Both monos and epis are closed under composition.
monic-∘: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ is-monic f
→ is-monic g
→ is-monic (f ∘ g)
= gm _ _ (fm _ _ (assoc _ _ _ ·· α ·· sym (assoc _ _ _)))
monic-∘ fm gm a b α
epic-∘: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ is-epic f
→ is-epic g
→ is-epic (f ∘ g)
= fe _ _ (ge _ _ (sym (assoc _ _ _) ·· α ·· (assoc _ _ _)))
epic-∘ fe ge a b α
_∘Mono_ : ∀ {a b c} → b ↪ c → a ↪ b → a ↪ c
(f ∘Mono g) .mor = f .mor ∘ g .mor
(f ∘Mono g) .monic = monic-∘ (f .monic) (g .monic)
_∘Epi_ : ∀ {a b c} → b ↠ c → a ↠ b → a ↠ c
(f ∘Epi g) .mor = f .mor ∘ g .mor
(f ∘Epi g) .epic = epic-∘ (f .epic) (g .epic)
If is monic, then must also be monic.
monic-cancell: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ is-monic (f ∘ g)
→ is-monic g
{f = f} fg-monic h h' p = fg-monic h h' $
monic-cancell (assoc _ _ _) ·· ap (f ∘_) p ·· assoc _ _ _ sym
Dually, if is epic, then must also be epic.
epic-cancelr: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ is-epic (f ∘ g)
→ is-epic f
{g = g} fg-epic h h' p = fg-epic h h' $
epic-cancelr _ _ _ ·· ap (_∘ g) p ·· sym (assoc _ _ _) assoc
Postcomposition with a mono is an embedding.
monic-postcomp-embedding: ∀ {a b c} {f : Hom b c}
→ is-monic f
→ is-embedding {A = Hom a b} (f ∘_)
=
monic-postcomp-embedding monic (Hom-set _ _) _ (monic _ _) injective→is-embedding
Likewise, precomposition with an epi is an embedding.
epic-precomp-embedding: ∀ {a b c} {f : Hom a b}
→ is-epic f
→ is-embedding {A = Hom b c} (_∘ f)
=
epic-precomp-embedding epic (Hom-set _ _) _ (epic _ _) injective→is-embedding
Sections🔗
A morphism is a section of another morphism when . The intuition for this name is that picks out a cross-section of from . For instance, could map animals to their species; a section of this map would have to pick out a canonical representative of each species from the set of all animals.
_section-of_ : (s : Hom b a) (r : Hom a b) → Type _
= r ∘ s ≡ id
s section-of r
section-of-is-prop: ∀ {s : Hom b a} {r : Hom a b}
→ is-prop (s section-of r)
= Hom-set _ _ _ _
section-of-is-prop
record has-section (r : Hom a b) : Type h where
constructor make-section
field
: Hom b a
section : section section-of r
is-section
open has-section public
: ∀ {a} → has-section (id {a})
id-has-section .section = id
id-has-section .is-section = idl _
id-has-section
section-of-∘: ∀ {a b c} {f : Hom c b} {g : Hom b c} {h : Hom b a} {i : Hom a b}
→ f section-of g → h section-of i
→ (h ∘ f) section-of (g ∘ i)
{f = f} {g = g} {h = h} {i = i} fg-sect hi-sect =
section-of-∘ (g ∘ i) ∘ h ∘ f ≡⟨ cat! C ⟩
(i ∘ h) ∘ f ≡⟨ ap (λ ϕ → g ∘ ϕ ∘ f) hi-sect ⟩
g ∘ (g ∘_) (idl _) ⟩
g ∘ id ∘ f ≡⟨ ap
g ∘ f ≡⟨ fg-sect ⟩
id ∎
section-∘: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ has-section f → has-section g → has-section (f ∘ g)
.section = g-sect .section ∘ f-sect .section
section-∘ f-sect g-sect {f = f} {g = g} f-sect g-sect .is-section =
section-∘ (f-sect .is-section) (g-sect .is-section) section-of-∘
If has a section, then is epic.
has-section→epic: ∀ {a b} {f : Hom a b}
→ has-section f
→ is-epic f
{f = f} f-sect g h p =
has-section→epic _ ⟩
g ≡˘⟨ idr (g ∘_) (f-sect .is-section) ⟩
g ∘ id ≡˘⟨ ap .section ≡⟨ assoc _ _ _ ⟩
g ∘ f ∘ f-sect (g ∘ f) ∘ f-sect .section ≡⟨ ap (_∘ f-sect .section) p ⟩
(h ∘ f) ∘ f-sect .section ≡˘⟨ assoc _ _ _ ⟩
.section ≡⟨ ap (h ∘_) (f-sect .is-section) ⟩
h ∘ f ∘ f-sect _ ⟩
h ∘ id ≡⟨ idr h ∎
Retracts🔗
A morphism is a retract of another morphism when . Note that this is the same equation involved in the definition of a section; retracts and sections always come in pairs. If sections solve a sort of “curation problem” where we are asked to pick out canonical representatives, then retracts solve a sort of “classification problem”.
_retract-of_ : (r : Hom a b) (s : Hom b a) → Type _
= r ∘ s ≡ id
r retract-of s
retract-of-is-prop: ∀ {s : Hom b a} {r : Hom a b}
→ is-prop (r retract-of s)
= Hom-set _ _ _ _
retract-of-is-prop
record has-retract (s : Hom b a) : Type h where
constructor make-retract
field
: Hom a b
retract : retract retract-of s
is-retract
open has-retract public
: ∀ {a} → has-retract (id {a})
id-has-retract .retract = id
id-has-retract .is-retract = idl _
id-has-retract
retract-of-∘: ∀ {a b c} {f : Hom c b} {g : Hom b c} {h : Hom b a} {i : Hom a b}
→ f retract-of g → h retract-of i
→ (h ∘ f) retract-of (g ∘ i)
= section-of-∘ hi-ret fg-ret
retract-of-∘ fg-ret hi-ret
retract-∘: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ has-retract f → has-retract g
→ has-retract (f ∘ g)
.retract = g-ret .retract ∘ f-ret .retract
retract-∘ f-ret g-ret .is-retract =
retract-∘ f-ret g-ret (f-ret .is-retract) (g-ret .is-retract) retract-of-∘
If has a retract, then is monic.
has-retract→monic: ∀ {a b} {f : Hom a b}
→ has-retract f
→ is-monic f
{f = f} f-ret g h p =
has-retract→monic _ ⟩
g ≡˘⟨ idl (_∘ g) (f-ret .is-retract) ⟩
id ∘ g ≡˘⟨ ap (f-ret .retract ∘ f) ∘ g ≡˘⟨ assoc _ _ _ ⟩
.retract ∘ (f ∘ g) ≡⟨ ap (f-ret .retract ∘_) p ⟩
f-ret .retract ∘ f ∘ h ≡⟨ assoc _ _ _ ⟩
f-ret (f-ret .retract ∘ f) ∘ h ≡⟨ ap (_∘ h) (f-ret .is-retract) ⟩
_ ⟩
id ∘ h ≡⟨ idl h ∎
A section that is also epic is a retract.
section-of+epic→retract-of: ∀ {a b} {s : Hom b a} {r : Hom a b}
→ s section-of r
→ is-epic s
→ s retract-of r
{s = s} {r = r} sect epic =
section-of+epic→retract-of (s ∘ r) id $
epic (s ∘ r) ∘ s ≡˘⟨ assoc s r s ⟩
(r ∘ s) ≡⟨ ap (s ∘_) sect ⟩
s ∘ _ ⟩
s ∘ id ≡⟨ idr _ ⟩
s ≡˘⟨ idl id ∘ s ∎
Dually, a retract that is also monic is a section.
retract-of+monic→section-of: ∀ {a b} {s : Hom b a} {r : Hom a b}
→ r retract-of s
→ is-monic r
→ r section-of s
{s = s} {r = r} ret monic =
retract-of+monic→section-of (s ∘ r) id $
monic
r ∘ s ∘ r ≡⟨ assoc r s r ⟩(r ∘ s) ∘ r ≡⟨ ap (_∘ r) ret ⟩
_ ⟩
id ∘ r ≡⟨ idl _ ⟩
r ≡˘⟨ idr r ∘ id ∎
Isos🔗
Maps and are inverses when we have and both equal to the identity. A map is invertible (or is an isomorphism) when there exists a for which and are inverses. An isomorphism is a choice of map , together with a specified inverse.
record Inverses (f : Hom a b) (g : Hom b a) : Type h where
field
: f ∘ g ≡ id
invl : g ∘ f ≡ id
invr
open Inverses
record is-invertible (f : Hom a b) : Type h where
field
: Hom b a
inv : Inverses f inv
inverses
open Inverses inverses public
: is-invertible inv
op .inv = f
op .inverses .Inverses.invl = invr inverses
op .inverses .Inverses.invr = invl inverses
op
record _≅_ (a b : Ob) : Type h where
field
to : Hom a b
: Hom b a
from : Inverses to from
inverses
open Inverses inverses public
open _≅_ public
A given map is invertible in at most one way: If you have with two inverses and , then not only are equal, but the witnesses of these equalities are irrelevant.
: ∀ {f : Hom a b} {g : Hom b a} → is-prop (Inverses f g)
Inverses-are-prop .Inverses.invl = Hom-set _ _ _ _ (x .invl) (y .invl) i
Inverses-are-prop x y i .Inverses.invr = Hom-set _ _ _ _ (x .invr) (y .invr) i
Inverses-are-prop x y i
: ∀ {f : Hom a b} → is-prop (is-invertible f)
is-invertible-is-prop {a = a} {b = b} {f = f} g h = p where
is-invertible-is-prop module g = is-invertible g
module h = is-invertible h
: g.inv ≡ h.inv
g≡h =
g≡h .inv ≡⟨ sym (idr _) ∙ ap₂ _∘_ refl (sym h.invl) ⟩
g.inv ∘ f ∘ h.inv ≡⟨ assoc _ _ _ ·· ap₂ _∘_ g.invr refl ·· idl _ ⟩
g.inv ∎
h
: g ≡ h
p .is-invertible.inv = g≡h i
p i .is-invertible.inverses =
p i (λ i → Inverses-are-prop {g = g≡h i}) g.inverses h.inverses i is-prop→pathp
We note that the identity morphism is always iso, and that isos compose:
: ∀ {a} → is-invertible (id {a})
id-invertible .is-invertible.inv = id
id-invertible .is-invertible.inverses .invl = idl id
id-invertible .is-invertible.inverses .invr = idl id
id-invertible
: a ≅ a
id-iso .to = id
id-iso .from = id
id-iso .inverses .invl = idl id
id-iso .inverses .invr = idl id
id-iso
= _≅_
Isomorphism
: {f : Hom a b} {f⁻¹ : Hom b a} {g : Hom b c} {g⁻¹ : Hom c b}
Inverses-∘ → Inverses f f⁻¹ → Inverses g g⁻¹ → Inverses (g ∘ f) (f⁻¹ ∘ g⁻¹)
{f = f} {f⁻¹} {g} {g⁻¹} finv ginv = record { invl = l ; invr = r } where
Inverses-∘ module finv = Inverses finv
module ginv = Inverses ginv
abstract
: (g ∘ f) ∘ f⁻¹ ∘ g⁻¹ ≡ id
l = (g ∘ f) ∘ f⁻¹ ∘ g⁻¹ ≡⟨ cat! C ⟩
l (f ∘ f⁻¹) ∘ g⁻¹ ≡⟨ (λ i → g ∘ finv.invl i ∘ g⁻¹) ⟩
g ∘
g ∘ id ∘ g⁻¹ ≡⟨ cat! C ⟩.invl ⟩
g ∘ g⁻¹ ≡⟨ ginv
id ∎
: (f⁻¹ ∘ g⁻¹) ∘ g ∘ f ≡ id
r = (f⁻¹ ∘ g⁻¹) ∘ g ∘ f ≡⟨ cat! C ⟩
r (g⁻¹ ∘ g) ∘ f ≡⟨ (λ i → f⁻¹ ∘ ginv.invr i ∘ f) ⟩
f⁻¹ ∘
f⁻¹ ∘ id ∘ f ≡⟨ cat! C ⟩.invr ⟩
f⁻¹ ∘ f ≡⟨ finv
id ∎
_∘Iso_ : a ≅ b → b ≅ c → a ≅ c
(f ∘Iso g) .to = g .to ∘ f .to
(f ∘Iso g) .from = f .from ∘ g .from
(f ∘Iso g) .inverses = Inverses-∘ (f .inverses) (g .inverses)
infixr 40 _∘Iso_
invertible-∘: ∀ {f : Hom b c} {g : Hom a b}
→ is-invertible f → is-invertible g
→ is-invertible (f ∘ g)
= record
invertible-∘ f-inv g-inv { inv = g-inv.inv ∘ f-inv.inv
; inverses = Inverses-∘ g-inv.inverses f-inv.inverses
}
where
module f-inv = is-invertible f-inv
module g-inv = is-invertible g-inv
_invertible⁻¹
: ∀ {f : Hom a b} → (f-inv : is-invertible f)
→ is-invertible (is-invertible.inv f-inv)
_invertible⁻¹ {f = f} f-inv .is-invertible.inv = f
_invertible⁻¹ f-inv .is-invertible.inverses .invl =
.invr f-inv
is-invertible_invertible⁻¹ f-inv .is-invertible.inverses .invr =
.invl f-inv
is-invertible
_Iso⁻¹ : a ≅ b → b ≅ a
(f Iso⁻¹) .to = f .from
(f Iso⁻¹) .from = f .to
(f Iso⁻¹) .inverses .invl = f .inverses .invr
(f Iso⁻¹) .inverses .invr = f .inverses .invl
We also note that invertible morphisms are both epic and monic.
: is-invertible f → is-monic f
invertible→monic {f = f} invert g h p =
invertible→monic
g ≡˘⟨ idl g ⟩(_∘ g) (is-invertible.invr invert) ⟩
id ∘ g ≡˘⟨ ap (inv ∘ f) ∘ g ≡˘⟨ assoc inv f g ⟩
(f ∘ g) ≡⟨ ap (inv ∘_) p ⟩
inv ∘ (f ∘ h) ≡⟨ assoc inv f h ⟩
inv ∘ (inv ∘ f) ∘ h ≡⟨ ap (_∘ h) (is-invertible.invr invert) ⟩
id ∘ h ≡⟨ idl h ⟩
h ∎where
open is-invertible invert
: is-invertible f → is-epic f
invertible→epic {f = f} invert g h p =
invertible→epic
g ≡˘⟨ idr g ⟩(g ∘_) (is-invertible.invl invert) ⟩
g ∘ id ≡˘⟨ ap (f ∘ inv) ≡⟨ assoc g f inv ⟩
g ∘ (g ∘ f) ∘ inv ≡⟨ ap (_∘ inv) p ⟩
(h ∘ f) ∘ inv ≡˘⟨ assoc h f inv ⟩
(f ∘ inv) ≡⟨ ap (h ∘_) (is-invertible.invl invert) ⟩
h ∘
h ∘ id ≡⟨ idr h ⟩
h ∎where
open is-invertible invert
: (f : a ≅ b) → is-monic (f .to)
iso→monic = invertible→monic (iso→invertible f)
iso→monic f
: (f : a ≅ b) → is-epic (f .to)
iso→epic = invertible→epic (iso→invertible f) iso→epic f
Furthermore, isomorphisms also yield section/retraction pairs.
inverses→to-has-section: ∀ {f : Hom a b} {g : Hom b a}
→ Inverses f g → has-section f
{g = g} inv .section = g
inverses→to-has-section .is-section = Inverses.invl inv
inverses→to-has-section inv
inverses→from-has-section: ∀ {f : Hom a b} {g : Hom b a}
→ Inverses f g → has-section g
{f = f} inv .section = f
inverses→from-has-section .is-section = Inverses.invr inv
inverses→from-has-section inv
inverses→to-has-retract: ∀ {f : Hom a b} {g : Hom b a}
→ Inverses f g → has-retract f
{g = g} inv .retract = g
inverses→to-has-retract .is-retract = Inverses.invr inv
inverses→to-has-retract inv
inverses→from-has-retract: ∀ {f : Hom a b} {g : Hom b a}
→ Inverses f g → has-retract g
{f = f} inv .retract = f
inverses→from-has-retract .is-retract = Inverses.invl inv
inverses→from-has-retract inv
: (f : a ≅ b) → has-section (f .to)
iso→to-has-section .section = f .from
iso→to-has-section f .is-section = f .invl
iso→to-has-section f
: (f : a ≅ b) → has-section (f .from)
iso→from-has-section .section = f .to
iso→from-has-section f .is-section = f .invr
iso→from-has-section f
: (f : a ≅ b) → has-retract (f .to)
iso→to-has-retract .retract = f .from
iso→to-has-retract f .is-retract = f .invr
iso→to-has-retract f
: (f : a ≅ b) → has-retract (f .from)
iso→from-has-retract .retract = f .to
iso→from-has-retract f .is-retract = f .invl iso→from-has-retract f
Using our lemmas about section/retraction pairs from before, we can show that if is a monic retract, then is an iso.
retract-of+monic→inverses: ∀ {a b} {s : Hom b a} {r : Hom a b}
→ r retract-of s
→ is-monic r
→ Inverses r s
.invl = ret
retract-of+monic→inverses ret monic .invr =
retract-of+monic→inverses ret monic retract-of+monic→section-of ret monic
We also have a dual result for sections and epis.
section-of+epic→inverses: ∀ {a b} {s : Hom b a} {r : Hom a b}
→ s retract-of r
→ is-epic r
→ Inverses r s
.invl =
section-of+epic→inverses sect epic
section-of+epic→retract-of sect epic.invr = sect section-of+epic→inverses sect epic
Hom-sets between isomorphic objects are equivalent. Crucially, this allows us to use univalence to transport properties between hom-sets.
iso→hom-equiv: ∀ {a a' b b'} → a ≅ a' → b ≅ b'
→ Hom a b ≃ Hom a' b'
= Iso→Equiv $
iso→hom-equiv f g (λ h → g .to ∘ h ∘ f .from) ,
(λ h → g .from ∘ h ∘ f .to )
iso (λ h →
.to ∘ (g .from ∘ h ∘ f .to) ∘ f .from ≡⟨ cat! C ⟩
g (g .to ∘ g .from) ∘ h ∘ (f .to ∘ f .from) ≡⟨ ap₂ (λ l r → l ∘ h ∘ r) (g .invl) (f .invl) ⟩
id ∘ h ∘ id ≡⟨ cat! C ⟩)
h ∎(λ h →
.from ∘ (g .to ∘ h ∘ f .from) ∘ f .to ≡⟨ cat! C ⟩
g (g .from ∘ g .to) ∘ h ∘ (f .from ∘ f .to) ≡⟨ ap₂ (λ l r → l ∘ h ∘ r) (g .invr) (f .invr) ⟩
id ∘ h ∘ id ≡⟨ cat! C ⟩) h ∎
If is invertible, then both pre and post-composition with are equivalences.
invertible-precomp-equiv: ∀ {a b c} {f : Hom a b}
→ is-invertible f
→ is-equiv {A = Hom b c} (_∘ f)
{f = f} f-inv = is-iso→is-equiv $
invertible-precomp-equiv (λ h → h ∘ f-inv.inv)
iso (λ h → sym (assoc _ _ _) ·· ap (h ∘_) f-inv.invr ·· idr h)
(λ h → sym (assoc _ _ _) ·· ap (h ∘_) f-inv.invl ·· idr h)
where module f-inv = is-invertible f-inv
invertible-postcomp-equiv: ∀ {a b c} {f : Hom b c}
→ is-invertible f
→ is-equiv {A = Hom a b} (f ∘_)
{f = f} f-inv = is-iso→is-equiv $
invertible-postcomp-equiv (λ h → f-inv.inv ∘ h)
iso (λ h → assoc _ _ _ ·· ap (_∘ h) f-inv.invl ·· idl h)
(λ h → assoc _ _ _ ·· ap (_∘ h) f-inv.invr ·· idl h)
where module f-inv = is-invertible f-inv