module Cat.Morphism {o h} (C : Precategory o h) where
open Precategory C public
private variable
: Ob
a b c d : Hom a b f
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
constructor make-mono
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
constructor make-epi
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
subst-is-monic: ∀ {a b} {f g : Hom a b}
→ f ≡ g
→ is-monic f
→ is-monic g
=
subst-is-monic f=g f-monic h i p (ap (_∘ h) f=g ·· p ·· ap (_∘ i) (sym f=g))
f-monic h i
subst-is-epic: ∀ {a b} {f g : Hom a b}
→ f ≡ g
→ is-epic f
→ is-epic g
=
subst-is-epic f=g f-epic h i p (ap (h ∘_) f=g ·· p ·· ap (i ∘_) (sym f=g)) f-epic h i
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
The identity map has a section (namely, itself), and the composite of maps with sections also has a section.
: ∀ {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-∘
Moreover, if has a section, then so does
section-cancell: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ has-section (f ∘ g)
→ has-section f
{g = g} s .section = g ∘ s .section
section-cancell {g = g} s .is-section = assoc _ _ _ ∙ s .is-section section-cancell
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 ∎
subst-section: ∀ {a b} {f g : Hom a b}
→ f ≡ g
→ has-section f
→ has-section g
.section = s .section
subst-section p s .is-section = ap₂ _∘_ (sym p) refl ∙ s .is-section subst-section p s
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
The identity map has a retract (namely, itself), and the composite of maps with retracts also has a retract.
: ∀ {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 so does
retract-cancelr: ∀ {a b c} {f : Hom b c} {g : Hom a b}
→ has-retract (f ∘ g)
→ has-retract g
{f = f} r .retract = r .retract ∘ f
retract-cancelr {f = f} r .is-retract = sym (assoc _ _ _) ∙ r .is-retract retract-cancelr
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 ∎
has-retract+epic→has-section: ∀ {a b} {f : Hom a b}
→ has-retract f
→ is-epic f
→ has-section f
.section = ret .retract
has-retract+epic→has-section ret epic .is-section =
has-retract+epic→has-section ret epic (ret .is-retract) epic
section-of+epic→retract-of
has-section+monic→has-retract: ∀ {a b} {f : Hom a b}
→ has-section f
→ is-monic f
→ has-retract f
.retract = sect .section
has-section+monic→has-retract sect monic .is-retract =
has-section+monic→has-retract sect monic (sect .is-section) monic retract-of+monic→section-of
Split monomorphisms🔗
A morphism is a split monomorphism if it merely has a retract.
: Hom a b → Type _
is-split-monic = ∥ has-retract f ∥ is-split-monic f
Every split mono is a mono, as being monic is a proposition.
: is-split-monic f → is-monic f
split-monic→mono = rec! has-retract→monic split-monic→mono
Split epimorphisms🔗
Dually, a morphism is a split epimorphism if it merely has a section.
: Hom a b → Type _
is-split-epic = ∥ has-section f ∥ is-split-epic f
Like split monos, every split epimorphism is an epimorphism.
: is-split-epic f → is-epic f
split-epic→epic = rec! has-section→epic split-epic→epic
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
: {a b c : Ob} {f : Hom b c} {f⁻¹ : Hom c b} {g : Hom a b} {g⁻¹ : Hom b a}
Inverses-∘ → Inverses f f⁻¹ → Inverses g g⁻¹ → Inverses (f ∘ g) (g⁻¹ ∘ f⁻¹)
{f = f} {f⁻¹} {g} {g⁻¹} finv ginv = record { invl = l ; invr = r } where
Inverses-∘ module finv = Inverses finv
module ginv = Inverses ginv
abstract
: (f ∘ g) ∘ g⁻¹ ∘ f⁻¹ ≡ id
l = (f ∘ g) ∘ g⁻¹ ∘ f⁻¹ ≡⟨ cat! C ⟩
l (g ∘ g⁻¹) ∘ f⁻¹ ≡⟨ (λ i → f ∘ ginv.invl i ∘ f⁻¹) ⟩
f ∘
f ∘ id ∘ f⁻¹ ≡⟨ cat! C ⟩.invl ⟩
f ∘ f⁻¹ ≡⟨ finv
id ∎
: (g⁻¹ ∘ f⁻¹) ∘ f ∘ g ≡ id
r = (g⁻¹ ∘ f⁻¹) ∘ f ∘ g ≡⟨ cat! C ⟩
r (f⁻¹ ∘ f) ∘ g ≡⟨ (λ i → g⁻¹ ∘ finv.invr i ∘ g) ⟩
g⁻¹ ∘
g⁻¹ ∘ id ∘ g ≡⟨ cat! C ⟩.invr ⟩
g⁻¹ ∘ g ≡⟨ ginv
id ∎
_∘Iso_ : ∀ {a b c : Ob} → b ≅ c → a ≅ b → a ≅ c
(f ∘Iso g) .to = f .to ∘ g .to
(f ∘Iso g) .from = g .from ∘ f .from
(f ∘Iso g) .inverses = Inverses-∘ (f .inverses) (g .inverses)
_∙Iso_ : ∀ {a b c : Ob} → a ≅ b → b ≅ c → a ≅ c
= g ∘Iso f
f ∙Iso g
infixr 40 _∘Iso_ _∙Iso_
infixr 41 _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-∘ f-inv.inverses g-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
: {f : Hom a b} {g : Hom b a} → f ∘ g ≡ id → g ∘ f ≡ id → Inverses f g
make-inverses .invl = p
make-inverses p q .invr = q
make-inverses p q
: {f : Hom a b} → (g : Hom b a) → f ∘ g ≡ id → g ∘ f ≡ id → is-invertible f
make-invertible .is-invertible.inv = g
make-invertible g p q .is-invertible.inverses .invl = p
make-invertible g p q .is-invertible.inverses .invr = q
make-invertible g p q
: (f : Hom a b) (g : Hom b a) → f ∘ g ≡ id → g ∘ f ≡ id → a ≅ b
make-iso ._≅_.to = f
make-iso f g p q ._≅_.from = g
make-iso f g p q ._≅_.inverses .Inverses.invl = p
make-iso f g p q ._≅_.inverses .Inverses.invr = q
make-iso f g p q
: ∀ {f : Hom a b} {g : Hom b a} → Inverses f g → is-invertible f
inverses→invertible .is-invertible.inv = _
inverses→invertible x .is-invertible.inverses = x
inverses→invertible x
: (f : Hom a b) → is-invertible f → a ≅ b
invertible→iso =
invertible→iso f x record
{ to = f
; from = x .is-invertible.inv
; inverses = x .is-invertible.inverses
}
is-invertible-inverse: {f : Hom a b} (g : is-invertible f) → is-invertible (g .is-invertible.inv)
=
is-invertible-inverse g record { inv = _ ; inverses = record { invl = invr g ; invr = invl g } }
where open Inverses (g .is-invertible.inverses)
: (i : a ≅ b) → is-invertible (i ._≅_.to)
iso→invertible = record { inv = i ._≅_.from ; inverses = i ._≅_.inverses }
iso→invertible i
private
≅-pathp-internal: (p : a ≡ c) (q : b ≡ d)
→ {f : a ≅ b} {g : c ≅ d}
→ PathP (λ i → Hom (p i) (q i)) (f ._≅_.to) (g ._≅_.to)
→ PathP (λ i → Hom (q i) (p i)) (f ._≅_.from) (g ._≅_.from)
→ PathP (λ i → p i ≅ q i) f g
.to = r i
≅-pathp-internal p q r s i .from = s i
≅-pathp-internal p q r s i {f} {g} r s i .inverses =
≅-pathp-internal p q (λ j → Inverses-are-prop {f = r j} {g = s j})
is-prop→pathp (f .inverses) (g .inverses) i
abstract
inverse-unique: {x y : Ob} (p : x ≡ y) {b d : Ob} (q : b ≡ d) (f : x ≅ b) (g : y ≅ d)
→ PathP (λ i → Hom (p i) (q i)) (f .to) (g .to)
→ PathP (λ i → Hom (q i) (p i)) (f .from) (g .from)
=
inverse-unique (λ a c p → ∀ {b d} (q : b ≡ d) (f : a ≅ b) (g : c ≅ d)
J' → PathP (λ i → Hom (p i) (q i)) (f .to) (g .to)
→ PathP (λ i → Hom (q i) (p i)) (f .from) (g .from))
λ x → J' (λ b d q → (f : x ≅ b) (g : x ≅ d)
→ PathP (λ i → Hom x (q i)) (f .to) (g .to)
→ PathP (λ i → Hom (q i) x) (f .from) (g .from))
λ y f g p →
.from ≡˘⟨ ap (f .from ∘_) (g .invl) ∙ idr _ ⟩
f .from ∘ g .to ∘ g .from ≡⟨ assoc _ _ _ ⟩
f (f .from ∘ g .to) ∘ g .from ≡⟨ ap (_∘ g .from) (ap (f .from ∘_) (sym p) ∙ f .invr) ∙ idl _ ⟩
.from ∎
g
≅-pathp: (p : a ≡ c) (q : b ≡ d) {f : a ≅ b} {g : c ≅ d}
→ PathP (λ i → Hom (p i) (q i)) (f ._≅_.to) (g ._≅_.to)
→ PathP (λ i → p i ≅ q i) f g
{f = f} {g = g} r = ≅-pathp-internal p q r (inverse-unique p q f g r)
≅-pathp p q
≅-pathp-from: (p : a ≡ c) (q : b ≡ d) {f : a ≅ b} {g : c ≅ d}
→ PathP (λ i → Hom (q i) (p i)) (f ._≅_.from) (g ._≅_.from)
→ PathP (λ i → p i ≅ q i) f g
{f = f} {g = g} r = ≅-pathp-internal p q (inverse-unique q p (f Iso⁻¹) (g Iso⁻¹) r) r
≅-pathp-from p q
: {f g : a ≅ b} → f ._≅_.to ≡ g ._≅_.to → f ≡ g
≅-path = ≅-pathp refl refl
≅-path
: {f g : a ≅ b} → f ._≅_.from ≡ g ._≅_.from → f ≡ g
≅-path-from = ≅-pathp-from refl refl
≅-path-from
: (∀ {a b} → is-contr (Hom a b)) → is-contr (a ≅ b)
≅-is-contr .centre =
≅-is-contr hom-contr (hom-contr .centre) (hom-contr .centre)
make-iso (is-contr→is-prop hom-contr _ _)
(is-contr→is-prop hom-contr _ _)
.paths f = ≅-path (hom-contr .paths (f .to))
≅-is-contr hom-contr
: (∀ {a b} → is-prop (Hom a b)) → is-prop (a ≅ b)
≅-is-prop = ≅-path (hom-prop (f .to) (g .to))
≅-is-prop hom-prop f g
↪-pathp: {a : I → Ob} {b : I → Ob} {f : a i0 ↪ b i0} {g : a i1 ↪ b i1}
→ PathP (λ i → Hom (a i) (b i)) (f .mor) (g .mor)
→ PathP (λ i → a i ↪ b i) f g
{a = a} {b} {f} {g} pa = go where
↪-pathp : PathP (λ i → a i ↪ b i) f g
go .mor = pa i
go i .monic {c = c} =
go i
is-prop→pathp(λ i → Π-is-hlevel {A = Hom c (a i)} 1
λ g → Π-is-hlevel {A = Hom c (a i)} 1
λ h → fun-is-hlevel {A = pa i ∘ g ≡ pa i ∘ h} 1
(Hom-set c (a i) g h))
(f .monic)
(g .monic)
i
↠-pathp: {a : I → Ob} {b : I → Ob} {f : a i0 ↠ b i0} {g : a i1 ↠ b i1}
→ PathP (λ i → Hom (a i) (b i)) (f .mor) (g .mor)
→ PathP (λ i → a i ↠ b i) f g
{a = a} {b} {f} {g} pa = go where
↠-pathp : PathP (λ i → a i ↠ b i) f g
go .mor = pa i
go i .epic {c = c} =
go i
is-prop→pathp(λ i → Π-is-hlevel {A = Hom (b i) c} 1
λ g → Π-is-hlevel {A = Hom (b i) c} 1
λ h → fun-is-hlevel {A = g ∘ pa i ≡ h ∘ pa i} 1
(Hom-set (b i) c g h))
(f .epic)
(g .epic)
i
subst-is-invertible: ∀ {x y} {f g : Hom x y}
→ f ≡ g
→ is-invertible f
→ is-invertible g
=
subst-is-invertible f=g f-inv .inv
make-invertible f(ap (_∘ f.inv) (sym f=g) ∙ f.invl)
(ap (f.inv ∘_) (sym f=g) ∙ f.invr)
where module f = is-invertible f-inv
Isomorphisms enjoy a 2-out-of-3 property: if any 2 of and are isomorphisms, then so is the third.
inverses-cancell: ∀ {f : Hom b c} {g : Hom a b} {g⁻¹ : Hom b a} {fg⁻¹ : Hom c a}
→ Inverses g g⁻¹ → Inverses (f ∘ g) fg⁻¹
→ Inverses f (g ∘ fg⁻¹)
inverses-cancelr: ∀ {f : Hom b c} {f⁻¹ : Hom c b} {g : Hom a b} {fg⁻¹ : Hom c a}
→ Inverses f f⁻¹ → Inverses (f ∘ g) fg⁻¹
→ Inverses g (fg⁻¹ ∘ f)
invertible-cancell: ∀ {f : Hom b c} {g : Hom a b}
→ is-invertible g → is-invertible (f ∘ g)
→ is-invertible f
invertible-cancelr: ∀ {f : Hom b c} {g : Hom a b}
→ is-invertible f → is-invertible (f ∘ g)
→ is-invertible g
We are early into our bootstrapping process for category theory, so the proofs of these lemmas are quite low-level, and thus omitted.
.invl =
inverses-cancell g-inv fg-inv _ _ _ ∙ fg-inv .invl
assoc .invr =
inverses-cancell g-inv fg-inv (idr _)
sym _∘_ refl (sym (g-inv .invl))
∙ ap₂ _ _ _
∙ assoc _∘_
∙ ap₂ (sym (assoc _ _ _)
(assoc _ _ _)
∙ sym _∘_ refl (fg-inv .invr)
∙ ap₂ _)
∙ idr
refl.invl
∙ g-inv
.invl =
inverses-cancelr f-inv fg-inv (idl _)
sym _∘_ (sym (f-inv .invr)) refl
∙ ap₂ (assoc _ _ _)
∙ sym _∘_ refl
∙ ap₂ (assoc _ _ _
_ _ _
∙ assoc _∘_ (fg-inv .invl) refl
∙ ap₂ _)
∙ idl .invr
∙ f-inv .invr =
inverses-cancelr f-inv fg-inv (assoc _ _ _) ∙ fg-inv .invr
sym
=
invertible-cancell g-inv fg-inv
inverses→invertible $
inverses-cancell(g-inv .is-invertible.inverses)
(fg-inv .is-invertible.inverses)
=
invertible-cancelr f-inv fg-inv
inverses→invertible $
inverses-cancelr(f-inv .is-invertible.inverses)
(fg-inv .is-invertible.inverses)
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
: is-invertible f → has-section f
invertible→to-has-section .section = is-invertible.inv f-inv
invertible→to-has-section f-inv .is-section = is-invertible.invl f-inv
invertible→to-has-section f-inv
: is-invertible f → has-retract f
invertible→to-has-retract .retract = is-invertible.inv f-inv
invertible→to-has-retract f-inv .is-retract = is-invertible.invr f-inv
invertible→to-has-retract f-inv
: is-invertible f → is-split-monic f
invertible→to-split-monic = pure (invertible→to-has-retract f-inv)
invertible→to-split-monic f-inv
: is-invertible f → is-split-epic f
invertible→to-split-epic = pure (invertible→to-has-section f-inv)
invertible→to-split-epic f-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
has-retract+epic→invertible: ∀ {a b} {f : Hom a b}
→ has-retract f
→ is-epic f
→ is-invertible f
.is-invertible.inv =
has-retract+epic→invertible f-ret epic .retract
f-ret .is-invertible.inverses =
has-retract+epic→invertible f-ret epic (f-ret .is-retract) epic
section-of+epic→inverses
has-section+monic→invertible: ∀ {a b} {f : Hom a b}
→ has-section f
→ is-monic f
→ is-invertible f
.is-invertible.inv =
has-section+monic→invertible f-sect monic .section
f-sect .is-invertible.inverses =
has-section+monic→invertible f-sect monic (f-sect .is-section) monic retract-of+monic→inverses
In fact, the mere existence of a retract of an epimorphism suffices to show that is invertible, as invertibility itself is a proposition. Put another way, every morphism that is both a split mono and an epi is invertible.
split-monic+epic→invertible: is-split-monic f
→ is-epic f
→ is-invertible f
=
split-monic+epic→invertible f-split-monic f-epic
∥-∥-rec is-invertible-is-prop(λ r → has-retract+epic→invertible r f-epic)
f-split-monic
A dual result holds for morphisms that are simultaneously split epic and monic.
split-epic+monic→invertible: is-split-epic f
→ is-monic f
→ is-invertible f
=
split-epic+monic→invertible f-split-epic f-monic
∥-∥-rec is-invertible-is-prop(λ s → has-section+monic→invertible s f-monic)
f-split-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 ∎
-- Optimized versions of Iso→Hom-equiv which yield nicer results
-- if only one isomorphism is needed.
dom-iso→hom-equiv: ∀ {a a' b} → a ≅ a'
→ Hom a b ≃ Hom a' b
= Iso→Equiv $
dom-iso→hom-equiv f (λ h → h ∘ f .from) ,
(λ h → h ∘ f .to )
iso (λ h →
(h ∘ f .to) ∘ f .from ≡⟨ sym (assoc _ _ _) ⟩
(f .to ∘ f .from) ≡⟨ ap (h ∘_) (f .invl) ⟩
h ∘ _ ⟩
h ∘ id ≡⟨ idr )
h ∎(λ h →
(h ∘ f .from) ∘ f .to ≡⟨ sym (assoc _ _ _) ⟩
(f .from ∘ f .to) ≡⟨ ap (h ∘_) (f .invr) ⟩
h ∘ _ ⟩
h ∘ id ≡⟨ idr )
h ∎
cod-iso→Hom-equiv: ∀ {a b b'} → b ≅ b'
→ Hom a b ≃ Hom a b'
= Iso→Equiv $
cod-iso→Hom-equiv f (λ h → f .to ∘ h) ,
(λ h → f .from ∘ h)
iso (λ h →
.to ∘ f .from ∘ h ≡⟨ assoc _ _ _ ⟩
f (f .to ∘ f .from) ∘ h ≡⟨ ap (_∘ h) (f .invl) ⟩
_ ⟩
id ∘ h ≡⟨ idl )
h ∎(λ h →
.from ∘ f .to ∘ h ≡⟨ assoc _ _ _ ⟩
f (f .from ∘ f .to) ∘ h ≡⟨ ap (_∘ h) (f .invr) ⟩
_ ⟩
id ∘ h ≡⟨ idl ) 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