Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 31db0bd

Browse files
committed
feat(category_theory/limits): add kernel pairs (#3925)
Add a subsingleton data structure expressing that a parallel pair of morphisms is a kernel pair of a given morphism. Another PR from my topos project. A pretty minimal API since I didn't need much there - I also didn't introduce anything like `has_kernel_pairs` largely because I didn't need it, but also because I don't know that it's useful for anyone, and it might conflict with ideas in the prop-limits branch.
1 parent a08fb2f commit 31db0bd

File tree

1 file changed

+170
-0
lines changed

1 file changed

+170
-0
lines changed
Lines changed: 170 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,170 @@
1+
/-
2+
Copyright (c) 2020 Bhavik Mehta. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Bhavik Mehta
5+
-/
6+
import category_theory.limits.shapes.equalizers
7+
import category_theory.limits.shapes.pullbacks
8+
import category_theory.limits.shapes.regular_mono
9+
10+
/-!
11+
# Kernel pairs
12+
13+
This file defines what it means for a parallel pair of morphisms `a b : R ⟶ X` to be the kernel pair
14+
for a morphism `f`.
15+
Some properties of kernel pairs are given, namely allowing one to transfer between
16+
the kernel pair of `f₁ ≫ f₂` to the kernel pair of `f₁`.
17+
It is also proved that if `f` is a coequalizer of some pair, and `a`,`b` is a kernel pair for `f` then
18+
it is a coequalizer of `a`,`b`.
19+
20+
## Implementation
21+
22+
The definition is essentially just a wrapper for `is_limit (pullback_cone.mk _ _ _)`, but the
23+
constructions given here are useful, yet awkward to present in that language, so a basic API
24+
is developed here.
25+
26+
## TODO
27+
28+
- Internal equivalence relations (or congruences) and the fact that every kernel pair induces one,
29+
and the converse in an effective regular category (WIP by b-mehta).
30+
31+
-/
32+
33+
universes v u u₂
34+
35+
namespace category_theory
36+
37+
open category_theory category_theory.category category_theory.limits
38+
39+
variables {C : Type u} [category.{v} C]
40+
41+
variables {R X Y Z : C} (f : X ⟶ Y) (a b : R ⟶ X)
42+
43+
/--
44+
`is_kernel_pair f a b` expresses that `(a, b)` is a kernel pair for `f`, i.e. `a ≫ f = b ≫ f`
45+
and the square
46+
R → X
47+
↓ ↓
48+
X → Y
49+
is a pullback square.
50+
This is essentially just a convenience wrapper over `is_limit (pullback_cone.mk _ _ _)`.
51+
-/
52+
structure is_kernel_pair :=
53+
(comm : a ≫ f = b ≫ f)
54+
(is_limit : is_limit (pullback_cone.mk _ _ comm))
55+
56+
attribute [reassoc] is_kernel_pair.comm
57+
58+
namespace is_kernel_pair
59+
60+
/-- The data expressing that `(a, b)` is a kernel pair is subsingleton. -/
61+
instance : subsingleton (is_kernel_pair f a b) :=
62+
⟨λ P Q, by { cases P, cases Q, congr, }⟩
63+
64+
/-- If `f` is a monomorphism, then `(𝟙 _, 𝟙 _)` is a kernel pair for `f`. -/
65+
def id_of_mono [mono f] : is_kernel_pair f (𝟙 _) (𝟙 _) :=
66+
{ comm := rfl,
67+
is_limit :=
68+
pullback_cone.is_limit_aux' _ $ λ s,
69+
begin
70+
refine ⟨s.snd, _, comp_id _, λ m m₁ m₂, _⟩,
71+
{ rw [← cancel_mono f, s.condition, pullback_cone.mk_fst, cancel_mono f],
72+
apply comp_id },
73+
rw [← m₂],
74+
apply (comp_id _).symm,
75+
end }
76+
77+
instance [mono f] : inhabited (is_kernel_pair f (𝟙 _) (𝟙 _)) := ⟨id_of_mono f⟩
78+
79+
variables {f a b}
80+
81+
/--
82+
Given a pair of morphisms `p`, `q` to `X` which factor through `f`, they factor through any kernel
83+
pair of `f`.
84+
-/
85+
def lift' {S : C} (k : is_kernel_pair f a b) (p q : S ⟶ X) (w : p ≫ f = q ≫ f) :
86+
{ t : S ⟶ R // t ≫ a = p ∧ t ≫ b = q } :=
87+
pullback_cone.is_limit.lift' k.is_limit _ _ w
88+
89+
/--
90+
If `(a,b)` is a kernel pair for `f₁ ≫ f₂` and `a ≫ f₁ = b ≫ f₁`, then `(a,b)` is a kernel pair for
91+
just `f₁`.
92+
That is, to show that `(a,b)` is a kernel pair for `f₁` it suffices to only show the square
93+
commutes, rather than to additionally show it's a pullback.
94+
-/
95+
def cancel_right {f₁ : X ⟶ Y} {f₂ : Y ⟶ Z} (comm : a ≫ f₁ = b ≫ f₁) (big_k : is_kernel_pair (f₁ ≫ f₂) a b) :
96+
is_kernel_pair f₁ a b :=
97+
{ comm := comm,
98+
is_limit := pullback_cone.is_limit_aux' _ $ λ s,
99+
begin
100+
let s' : pullback_cone (f₁ ≫ f₂) (f₁ ≫ f₂) := pullback_cone.mk s.fst s.snd (s.condition_assoc _),
101+
refine ⟨big_k.is_limit.lift s',
102+
big_k.is_limit.fac _ walking_cospan.left,
103+
big_k.is_limit.fac _ walking_cospan.right,
104+
λ m m₁ m₂, _⟩,
105+
apply big_k.is_limit.hom_ext,
106+
refine ((pullback_cone.mk a b _) : pullback_cone (f₁ ≫ f₂) _).equalizer_ext _ _,
107+
apply m₁.trans (big_k.is_limit.fac s' walking_cospan.left).symm,
108+
apply m₂.trans (big_k.is_limit.fac s' walking_cospan.right).symm,
109+
end }
110+
111+
/--
112+
If `(a,b)` is a kernel pair for `f₁ ≫ f₂` and `f₂` is mono, then `(a,b)` is a kernel pair for
113+
just `f₁`.
114+
The converse of `comp_of_mono`.
115+
-/
116+
def cancel_right_of_mono {f₁ : X ⟶ Y} {f₂ : Y ⟶ Z} [mono f₂] (big_k : is_kernel_pair (f₁ ≫ f₂) a b) :
117+
is_kernel_pair f₁ a b :=
118+
cancel_right (begin rw [← cancel_mono f₂, assoc, assoc, big_k.comm] end) big_k
119+
120+
/--
121+
If `(a,b)` is a kernel pair for `f₁` and `f₂` is mono, then `(a,b)` is a kernel pair for `f₁ ≫ f₂`.
122+
The converse of `cancel_right_of_mono`.
123+
-/
124+
def comp_of_mono {f₁ : X ⟶ Y} {f₂ : Y ⟶ Z} [mono f₂] (small_k : is_kernel_pair f₁ a b) :
125+
is_kernel_pair (f₁ ≫ f₂) a b :=
126+
{ comm := by rw [small_k.comm_assoc],
127+
is_limit := pullback_cone.is_limit_aux' _ $ λ s,
128+
begin
129+
refine ⟨_, _, _, _⟩,
130+
apply (pullback_cone.is_limit.lift' small_k.is_limit s.fst s.snd _).1,
131+
rw [← cancel_mono f₂, assoc, s.condition, assoc],
132+
apply (pullback_cone.is_limit.lift' small_k.is_limit s.fst s.snd _).2.1,
133+
apply (pullback_cone.is_limit.lift' small_k.is_limit s.fst s.snd _).2.2,
134+
intros m m₁ m₂,
135+
apply small_k.is_limit.hom_ext,
136+
refine ((pullback_cone.mk a b _) : pullback_cone f₁ _).equalizer_ext _ _,
137+
rwa (pullback_cone.is_limit.lift' small_k.is_limit s.fst s.snd _).2.1,
138+
rwa (pullback_cone.is_limit.lift' small_k.is_limit s.fst s.snd _).2.2,
139+
end }
140+
141+
/--
142+
If `(a,b)` is the kernel pair of `f`, and `f` is a coequalizer morphism for some parallel pair, then
143+
`f` is a coequalizer morphism of `a` and `b`.
144+
-/
145+
def to_coequalizer (k : is_kernel_pair f a b) [r : regular_epi f] :
146+
is_colimit (cofork.of_π f k.comm) :=
147+
begin
148+
let t := k.is_limit.lift (pullback_cone.mk _ _ r.w),
149+
have ht : t ≫ a = r.left := k.is_limit.fac _ walking_cospan.left,
150+
have kt : t ≫ b = r.right := k.is_limit.fac _ walking_cospan.right,
151+
apply cofork.is_colimit.mk _ _ _ _,
152+
{ intro s,
153+
apply (cofork.is_colimit.desc' r.is_colimit s.π _).1,
154+
rw [← ht, assoc, s.condition, reassoc_of kt] },
155+
{ intro s,
156+
apply (cofork.is_colimit.desc' r.is_colimit s.π _).2 },
157+
{ intros s m w,
158+
apply r.is_colimit.hom_ext,
159+
rintro ⟨⟩,
160+
change (r.left ≫ f) ≫ m = (r.left ≫ f) ≫ _,
161+
rw [assoc, assoc],
162+
congr' 1,
163+
erw (cofork.is_colimit.desc' r.is_colimit s.π _).2,
164+
apply w walking_parallel_pair.one,
165+
erw (cofork.is_colimit.desc' r.is_colimit s.π _).2,
166+
apply w walking_parallel_pair.one }
167+
end
168+
169+
end is_kernel_pair
170+
end category_theory

0 commit comments

Comments
 (0)