Kernel pairs
This file defines what it means for a parallel pair of morphisms a b : R ⟶ X
to be the kernel pair
for a morphism f
.
Some properties of kernel pairs are given, namely allowing one to transfer between
the kernel pair of f₁ ≫ f₂
to the kernel pair of f₁
.
It is also proved that if f
is a coequalizer of some pair, and a
,b
is a kernel pair for f
then
it is a coequalizer of a
,b
.
Implementation
The definition is essentially just a wrapper for is_limit (pullback_cone.mk _ _ _)
, but the
constructions given here are useful, yet awkward to present in that language, so a basic API
is developed here.
TODO
- Internal equivalence relations (or congruences) and the fact that every kernel pair induces one, and the converse in an effective regular category (WIP by b-mehta).
- comm : a ≫ f = b ≫ f
- is_limit : category_theory.limits.is_limit (category_theory.limits.pullback_cone.mk a b _)
is_kernel_pair f a b
expresses that (a, b)
is a kernel pair for f
, i.e. a ≫ f = b ≫ f
and the square
R → X
↓ ↓
X → Y
is a pullback square.
This is essentially just a convenience wrapper over is_limit (pullback_cone.mk _ _ _)
.
The data expressing that (a, b)
is a kernel pair is subsingleton.
If f
is a monomorphism, then (𝟙 _, 𝟙 _)
is a kernel pair for f
.
Equations
- category_theory.is_kernel_pair.id_of_mono f = {comm := _, is_limit := (category_theory.limits.pullback_cone.mk (𝟙 X) (𝟙 X) _).is_limit_aux' (λ (s : category_theory.limits.pullback_cone f f), ⟨s.snd, _⟩)}
Equations
Given a pair of morphisms p
, q
to X
which factor through f
, they factor through any kernel
pair of f
.
Equations
- k.lift' p q w = category_theory.limits.pullback_cone.is_limit.lift' k.is_limit p q w
If (a,b)
is a kernel pair for f₁ ≫ f₂
and a ≫ f₁ = b ≫ f₁
, then (a,b)
is a kernel pair for
just f₁
.
That is, to show that (a,b)
is a kernel pair for f₁
it suffices to only show the square
commutes, rather than to additionally show it's a pullback.
Equations
- category_theory.is_kernel_pair.cancel_right comm big_k = {comm := comm, is_limit := (category_theory.limits.pullback_cone.mk a b comm).is_limit_aux' (λ (s : category_theory.limits.pullback_cone f₁ f₁), let s' : category_theory.limits.pullback_cone (f₁ ≫ f₂) (f₁ ≫ f₂) := category_theory.limits.pullback_cone.mk s.fst s.snd _ in ⟨big_k.is_limit.lift s', _⟩)}
If (a,b)
is a kernel pair for f₁ ≫ f₂
and f₂
is mono, then (a,b)
is a kernel pair for
just f₁
.
The converse of comp_of_mono
.
Equations
- big_k.cancel_right_of_mono = category_theory.is_kernel_pair.cancel_right _ big_k
If (a,b)
is a kernel pair for f₁
and f₂
is mono, then (a,b)
is a kernel pair for f₁ ≫ f₂
.
The converse of cancel_right_of_mono
.
Equations
- small_k.comp_of_mono = {comm := _, is_limit := (category_theory.limits.pullback_cone.mk a b _).is_limit_aux' (λ (s : category_theory.limits.pullback_cone (f₁ ≫ f₂) (f₁ ≫ f₂)), ⟨(category_theory.limits.pullback_cone.is_limit.lift' small_k.is_limit s.fst s.snd _).val, _⟩)}
If (a,b)
is the kernel pair of f
, and f
is a coequalizer morphism for some parallel pair, then
f
is a coequalizer morphism of a
and b
.
Equations
- k.to_coequalizer = let t : (category_theory.limits.pullback_cone.mk category_theory.regular_epi.left category_theory.regular_epi.right category_theory.regular_epi.w).X ⟶ (category_theory.limits.pullback_cone.mk a b _).X := k.is_limit.lift (category_theory.limits.pullback_cone.mk category_theory.regular_epi.left category_theory.regular_epi.right category_theory.regular_epi.w) in category_theory.limits.cofork.is_colimit.mk (category_theory.limits.cofork.of_π f _) (λ (s : category_theory.limits.cofork a b), (category_theory.limits.cofork.is_colimit.desc' category_theory.regular_epi.is_colimit s.π _).val) _ _