The snake lemma in terms of modules #
The snake lemma is proven in Algebra/Homology/ShortComplex/SnakeLemma.lean
for all abelian
categories, but for definitional equality and universe issues we reprove them here for modules.
Main results #
SnakeLemma.δ
: The connecting homomorphism guranteed by the snake lemma.SnakeLemma.exact_δ_left
: The connecting homomorphism is exact on the right.SnakeLemma.exact_δ_right
: The connecting homomorphism is exact on the left.
Suppose we have an exact commutative diagram
K₂ -F-→ K₃
| |
ι₂ ι₃
↓ ↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
| |
π₁ π₂
↓ ↓
C₁ -G-→ C₂
such that f₂
is surjective with a (set-theoretic) section σ
, g₁
is injective with a
(set-theoretic) retraction ρ
, and that ι₃
is injective and π₁
is surjective.
Snake Lemma Supppose we have an exact commutative diagram
K₃
|
ι₃
↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
|
π₁
↓
C₁
such that f₂
is surjective with a (set-theoretic) section σ
, g₁
is injective with a
(set-theoretic) retraction ρ
,
then the map π₁ ∘ ρ ∘ i₂ ∘ σ ∘ ι₃
is a linear map from K₃
to C₁
.
Also see SnakeLemma.δ'
for a noncomputable version
that does not require an explicit section and retraction.
Equations
- SnakeLemma.δ i₁ i₂ i₃ f₁ f₂ hf g₁ g₂ hg h₁ h₂ σ hσ ρ hρ ι₃ hι₃ π₁ hπ₁ = { toFun := fun (x : K₃) => π₁ (ρ (i₂ (σ (ι₃ x)))), map_add' := ⋯, map_smul' := ⋯ }
Instances For
Supppose we have an exact commutative diagram
K₂ -F-→ K₃
| |
ι₂ ι₃
↓ ↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
|
π₁
↓
C₁
such that f₂
is surjective with a (set-theoretic) section σ
, g₁
is injective with a
(set-theoretic) retraction ρ
, and ι₃
is injective, then K₂ -F→ K₂ -δ→ C₁
is exact.
Supppose we have an exact commutative diagram
K₃
|
ι₃
↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
| |
π₁ π₂
↓ ↓
C₁ -G-→ C₂
such that f₂
is surjective with a (set-theoretic) section σ
, g₁
is injective with a
(set-theoretic) retraction ρ
, and π₁
is surjective, then K₂ -δ→ C₁ -G→ C₂
is exact.
Supppose we have an exact commutative diagram
K₃
|
ι₃
↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
|
π₁
↓
C₁
such that f₂
is surjective and g₁
is injective,
then this is the linear map K₃ → C₁
given by the snake lemma.
Also see SnakeLemma.δ
for a computable version.
Equations
- SnakeLemma.δ' i₁ i₂ i₃ f₁ f₂ hf g₁ g₂ hg h₁ h₂ ι₃ hι₃ π₁ hπ₁ hf₂ hg₁ = SnakeLemma.δ i₁ i₂ i₃ f₁ f₂ hf g₁ g₂ hg h₁ h₂ (Function.surjInv hf₂) ⋯ (Function.invFun ⇑g₁) ⋯ ι₃ hι₃ π₁ hπ₁
Instances For
Supppose we have an exact commutative diagram
K₂ -F-→ K₃
| |
ι₂ ι₃
↓ ↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
|
π₁
↓
C₁
such that f₂
is surjective, g₁
is injective, and ι₃
is injective,
then K₂ -F→ K₂ -δ→ C₁
is exact.
Supppose we have an exact commutative diagram
K₃
|
ι₃
↓
M₁ -f₁→ M₂ -f₂→ M₃
| | |
i₁ i₂ i₃
↓ ↓ ↓
N₁ -g₁→ N₂ -g₂→ N₃
| |
π₁ π₂
↓ ↓
C₁ -G-→ C₂
such that f₂
is surjective, g₁
is injective, and π₁
is surjective,
then K₂ -δ→ C₁ -G→ C₂
is exact.