module Cat.Morphism.Strong.Epi {o ℓ} (C : Precategory o ℓ) where
open Cat.Reasoning C
Strong epimorphisms🔗
A strong epimorphism is an epimorphism which is, additionally, left orthogonal to every monomorphism. Unfolding that definition, for to be a strong epimorphism means that, given any mono, and arbitrarily fit into a commutative diagram like
there is a contractible space of maps which make the two triangles commute. In the particular case of strong epimorphisms, it actually suffices to have any lift: they are automatically unique.
: ∀ {a b} → Hom a b → Type _
is-strong-epi = is-epic f × ∀ {c d} (m : c ↪ d) → m⊥m C f (m .mor)
is-strong-epi f
lifts→is-strong-epi: ∀ {a b} {f : Hom a b}
→ is-epic f
→ ( ∀ {c d} (m : c ↪ d) {u} {v} → v ∘ f ≡ m .mor ∘ u
→ Lifting C f (m .mor) u v)
→ is-strong-epi f
= epic , λ {c} {d} mm sq →
lifts→is-strong-epi epic lift-it (lift-it mm sq) λ { (x , p , q) → Σ-prop-path!
contr (mm .monic _ _ (sym (q ∙ sym (lift-it mm sq .snd .snd)))) }
abstract
is-strong-epi-is-prop: ∀ {a b} (f : Hom a b) → is-prop (is-strong-epi f)
= hlevel 1 is-strong-epi-is-prop f
To see that the uniqueness needed for orthogonality against a monomorphism is redundant, suppose you’d had two fillers as in
Since is a monomorphism, it suffices to have but by commutativity of the lower triangles we have
Properties🔗
The first thing we show is that strong epimorphisms are closed under
composition. It will suffice to show that the composite
of a pair of strong epimorphisms is epic, and that it lifts against
every monomorphism. But this is just a pasting of existing results: we
already know that being epic is closed under composition
, and that so is
lifting against a given map
— in
this case, an arbitrary monomorphism
strong-epi-∘: ∀ {a b c} (f : Hom b c) (g : Hom a b)
→ is-strong-epi f
→ is-strong-epi g
→ is-strong-epi (f ∘ g)
(f-epi , f-str) (g-epi , g-str) =
strong-epi-∘ f g (epic-∘ f-epi g-epi) λ e → ∘l-lifts-against C
lifts→is-strong-epi (orthogonal→lifts-against C (f-str e))
(orthogonal→lifts-against C (g-str e))
Additionally, there is a partial converse to this result: If the composite is a strong epi, then is, too! It’ll help to diagram the situation:
We can now name our assumptions: we’re given to lift against We don’t have a as before, but we can take to get a lift
strong-epi-cancelr: ∀ {a b c} (f : Hom b c) (g : Hom a b)
→ is-strong-epi (f ∘ g)
→ is-strong-epi f
(fg-epi , fg-str) =
strong-epi-cancelr f g
lifts→is-strong-epi f-epi f-strwhere
: is-epic f
f-epi = epic-cancelr fg-epi
f-epi
: ∀ {c d} (m : c ↪ d) {u} {v} → v ∘ f ≡ m .mor ∘ u → _
f-str {u} {v} vf=mu =
f-str m let (w , wfg=ug , mw=v) = fg-str m (extendl vf=mu) .centre
in w , m .monic (w ∘ f) u (pulll mw=v ∙ vf=mu) , mw=v
As an immediate consequence of the definition, a monic strong epi is an isomorphism. This is because, being left orthogonal to all monos, it’d be, in particular, left orthogonal to itself, and the only self-orthogonal maps are isos.
strong-epi+mono→invertible: ∀ {a b} {f : Hom a b} → is-strong-epi f → is-monic f → is-invertible f
(_ , strong) mono =
strong-epi+mono→invertible _ (strong (record { monic = mono })) self-orthogonal→invertible C
Regular epis are strong🔗
Suppose that is a regular epimorphism, that is, it identifies as some quotient of — the stuff of is that of but with new potential identifications thrown in. Since we’re taking “strong epimorphism” as the definition of “well-behaved epimorphism”, we’d certainly like regular epis to be strong epis!
This is fortunately the case. Suppose that is the coequaliser of some maps 1, and that is a monomorphism we want to lift against.
By the universal property of a coequaliser, to “slide over” to a map it suffices to show that it also coequalises the pair i.e. that Since is a mono, it’ll suffice to show that but note that since the square commutes. Then we have
so there is a map such that — that’s commutativity of the top triangle handled. For the bottom triangle, since is a regular epic (thus an epic), to show it’ll suffice to show that But by assumption, and by the universal property! We’re done.
is-regular-epi→is-strong-epi: ∀ {a b} (f : Hom a b)
→ is-regular-epi C f
→ is-strong-epi f
{a} {b} f regular =
is-regular-epi→is-strong-epi
lifts→is-strong-epi.is-regular-epi→is-epic
r(λ m x → map m x , r.factors , lemma m x)
where
module r = is-regular-epi regular renaming (arr₁ to s ; arr₂ to t)
module _ {c} {d} (z : c ↪ d) {u} {v} (vf=zu : v ∘ f ≡ z .mor ∘ u) where
module z = _↪_ z
: Hom b c
map = r.universal {e' = u} $ z.monic _ _ $
map .mor ∘ u ∘ r.s ≡⟨ extendl (sym vf=zu) ⟩
z .s ≡⟨ refl⟩∘⟨ r.coequal ⟩
v ∘ f ∘ r.t ≡˘⟨ extendl (sym vf=zu) ⟩
v ∘ f ∘ r.mor ∘ u ∘ r.t ∎
z = r.is-regular-epi→is-epic _ _ $
lemma (vf=zu ∙ pushr (sym r.factors)) sym
Images🔗
Now we come to the raison d’être for strong epimorphisms: Images. The definition of image we use is very complicated, and the justification is already present there, but the short of it is that the image of a morphism is a monomorphism which is universal amongst those through which factors.
Since images have a universal property, and one involving comma categories of slice categories at that, they are tricky to come by. However, in the case where we factor as
and the epimorphism is strong, then we automatically have an image factorisation of on our hands!
strong-epi-mono→image: ∀ {a b im} (f : Hom a b)
→ (a→im : Hom a im) → is-strong-epi a→im
→ (im→b : Hom im b) → is-monic im→b
→ im→b ∘ a→im ≡ f
→ Image C f
(_ , str-epi) im→b mono fact = go where
strong-epi-mono→image f a→im open Initial
open /-Obj
open /-Hom
open ↓Obj
open ↓Hom
: ↓Obj (!Const (cut f)) (Forget-full-subcat {P = is-monic ⊙ map})
obj .x = tt
obj .y = cut im→b , mono
obj .map = record { map = a→im ; commutes = fact } obj
Actually, for an image factorisation, we don’t need that be an epimorphism — we just need it to be orthogonal to every monomorphism. This turns out to be precisely the data of being initial in the relevant comma categories.
: Image C f
go .bot = obj
go .has⊥ other = contr dh unique where
go module o = ↓Obj other
=
the-lifting
str-epi(record { monic = o.y .snd })
{u = o.map .map}
{v = im→b} (sym (o.map .commutes ∙ sym fact))
: ↓Hom (!Const (cut f)) _ obj other
dh .α = tt
dh .β .map = the-lifting .centre .fst
dh .β .commutes = the-lifting .centre .snd .snd
dh .sq = ext (idr _ ∙ sym (the-lifting .centre .snd .fst))
dh
: ∀ om → dh ≡ om
unique = ↓Hom-path _ _ refl $ ext $ ap fst $ the-lifting .paths $
unique om .β .map , sym (ap map (om .sq)) ∙ idr _ , om .β .commutes om
In the lex case🔗
Suppose that is additionally left exact, or more restrictively, that it has all equalisers. In that case, a map left orthogonal to all monomorphisms is automatically an epimorphism, thus a strong epi. Let’s see how. First, there’s a quick observation to be made about epimorphisms: if is such that there exists a with then is an epimorphism. You can think of this as a special case of “ epic implies epic” or as a short calculation:
retract-is-epi: ∀ {a b} {f : Hom a b} {g : Hom b a}
→ f ∘ g ≡ id
→ is-epic f
{f = f} {g} p α β q =
retract-is-epi
α ≡⟨ intror p ⟩
α ∘ f ∘ g ≡⟨ extendl q ⟩
β ∘ f ∘ g ≡⟨ elimr p ⟩ β ∎
We already know that if lifts exist and the map is epic, then it’s a strong epi, so let’s assume that lifts exist — we’ll have no need for uniqueness, here. Given and to lift against, form their equaliser and arrange them like
equaliser-lifts→is-strong-epi: ∀ {a b} {f : Hom a b}
→ (∀ {a b} (f g : Hom a b) → Equaliser C f g)
→ ( ∀ {c d} (m : c ↪ d) {u} {v} → v ∘ f ≡ m .mor ∘ u
→ Σ[ w ∈ Hom b c ] ((w ∘ f ≡ u) × (m .mor ∘ w ≡ v)))
→ is-strong-epi f
{f = f} eqs ls = lifts→is-strong-epi epi ls where equaliser-lifts→is-strong-epi
By the universal property of since there’s there’s a unique map such that Arranging and the identity(!) into a lifting square like the one above, we conclude the existence of a dotted map satisfying, most importantly, — so that being a retract, is an epimorphism.
: is-epic f
epi =
epi u v uf=vf let
module ker = Equaliser (eqs u v)
= ker.universal uf=vf
k (w , p , q) = ls
(record { monic = is-equaliser→is-monic _ ker.has-is-eq })
{u = k} {v = id}
(idl _ ∙ sym ker.factors)
: is-epic ker.equ
e-epi = retract-is-epi q e-epi
Now, is the universal map which equalises and — so that we have and since we’ve just shown that is epic, this means we have — exactly what we wanted!
in e-epi u v ker.equal
Extremal epimorphisms🔗
Another well-behaved subclass of epimorphism are the extremal epimorphisms: An epimorphism is extremal if when, given a factorisation through a monomorphism then is an isomorphism. In a finitely complete category, every extremal epimorphism is strong; the converse is immediate.
: ∀ {a b} → Hom a b → Type _
is-extremal-epi {a} {b} e =
is-extremal-epi ∀ {c} (m : c ↪ b) (g : Hom a c)
→ e ≡ m .mor ∘ g
→ is-invertible (m .mor)
is-extremal-epi→is-strong-epi: ∀ {a b} {e : Hom a b}
→ Finitely-complete C
→ is-extremal-epi e
→ is-strong-epi e
{a} {b} {e} lex extremal =
is-extremal-epi→is-strong-epi .equalisers λ w → Mk.the-lift w where
equaliser-lifts→is-strong-epi lexmodule lex = Finitely-complete lex
We adapt the proof from (Borceux
1994, vol. 1, sec. 4.3.7). After equaliser-lifts→is-strong-epi
, it will
suffice to construct some lift for a square with
with
monic. Pull
back along
to obtain the square
and obtain the unique factorisation Note that the map is a monomorphism since it results from pulling back a monomorphism.
module Mk {c d : Ob} (m : c ↪ d) {u : Hom a c} {v : Hom b d}
(wit : v ∘ e ≡ m .mor ∘ u) where
module P = Pullback (lex.pullbacks v (m .mor)) renaming (p₁ to q ; p₂ to p)
: Hom a P.apex
r = P.universal {p₁' = e} {p₂' = u} wit
r
abstract
: is-monic P.q
q-mono = is-monic→pullback-is-monic (m .monic) (rotate-pullback P.has-is-pb) q-mono
We thus have a factorisation of through a monomorphism which since was assumed extremal, must be an isomorphism. We define the diagonal map to be and compute that it commutes appropriately:
: is-invertible P.q
q-iso = extremal record{ monic = q-mono } r (sym P.p₁∘universal)
q-iso
= q-iso .is-invertible.inv
q⁻¹
: Σ (Hom b c) λ w → (w ∘ e ≡ u) × (m .mor ∘ w ≡ v)
the-lift .fst = P.p ∘ q⁻¹
the-lift .snd .fst = m .monic _ _ $
the-lift .mor ∘ (P.p ∘ q⁻¹) ∘ e ≡⟨ extendl (pulll (sym P.square)) ⟩
m (v ∘ P.q) ∘ q⁻¹ ∘ e ≡⟨ cancel-inner (q-iso .is-invertible.invl) ⟩
v ∘ e ≡⟨ wit ⟩.mor ∘ u ∎
m .snd .snd = invertible→epic q-iso _ _ $
the-lift (m .mor ∘ (P.p ∘ q⁻¹)) ∘ P.q ≡⟨ pullr (cancelr (q-iso .is-invertible.invr)) ⟩
.mor ∘ P.p ≡˘⟨ P.square ⟩
m .q ∎ v ∘ P
is-strong-epi→is-extremal-epi: ∀ {a b} {e : Hom a b}
→ is-strong-epi e
→ ∀ {c} (m : c ↪ b) (g : Hom a c) → e ≡ m .mor ∘ g → is-invertible (m .mor)
(s , ortho) m g p =
is-strong-epi→is-extremal-epi (inv' .centre .fst) (inv' .centre .snd .snd)
make-invertible (m .monic _ _ (pulll (inv' .centre .snd .snd) ∙ id-comm-sym))
where
= ortho m (idl _ ∙ p) inv'
invertible→strong-epi: ∀ {a b} {f : Hom a b}
→ is-invertible f
→ is-strong-epi f
=
invertible→strong-epi f-inv λ m →
invertible→epic f-inv , (m .mor) f-inv
invertible→left-orthogonal C
subst-is-strong-epi: ∀ {a b} {f g : Hom a b}
→ f ≡ g
→ is-strong-epi f
→ is-strong-epi g
=
subst-is-strong-epi f=g f-strong-epi (subst-is-epic f=g (f-strong-epi .fst)) λ m vg=mu →
lifts→is-strong-epi let (h , hf=u , mh=v) = f-strong-epi .snd m (ap₂ _∘_ refl f=g ∙ vg=mu) .centre
in h , ap (h ∘_) (sym f=g) ∙ hf=u , mh=v
If you care, is for “relation” — the intuition is that specifies the relations imposed on to get ↩︎
References
- Borceux, Francis. 1994. Handbook of Categorical Algebra. Vol. 1. Encyclopedia of Mathematics and Its Applications. Cambridge University Press. https://doi.org/10.1017/CBO9780511525858.