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

Commit bd7d034

Browse files
committed
feat(ring_theory/nilpotent): add lemma module.End.is_nilpotent_mapq (#11831)
Together with the other lemmas necessary for its proof.
1 parent b905eb6 commit bd7d034

File tree

3 files changed

+51
-0
lines changed

3 files changed

+51
-0
lines changed

src/linear_algebra/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -633,6 +633,14 @@ rfl
633633
lemma comap_mono {f : M →ₛₗ[σ₁₂] M₂} {q q' : submodule R₂ M₂} :
634634
q ≤ q' → comap f q ≤ comap f q' := preimage_mono
635635

636+
lemma le_comap_pow_of_le_comap (p : submodule R M) {f : M →ₗ[R] M} (h : p ≤ p.comap f) (k : ℕ) :
637+
p ≤ p.comap (f^k) :=
638+
begin
639+
induction k with k ih,
640+
{ simp [linear_map.one_eq_id], },
641+
{ simp [linear_map.iterate_succ, comap_comp, h.trans (comap_mono ih)], },
642+
end
643+
636644
section
637645
variables [ring_hom_surjective σ₁₂]
638646

src/linear_algebra/quotient.lean

Lines changed: 29 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -264,6 +264,35 @@ p.liftq (q.mkq.comp f) $ by simpa [ker_comp] using h
264264
theorem mapq_mkq (f : M →ₛₗ[τ₁₂] M₂) {h} : (mapq p q f h).comp p.mkq = q.mkq.comp f :=
265265
by ext x; refl
266266

267+
@[simp] lemma mapq_zero (h : p ≤ q.comap (0 : M →ₛₗ[τ₁₂] M₂) := by simp) :
268+
p.mapq q (0 : M →ₛₗ[τ₁₂] M₂) h = 0 :=
269+
by { ext, simp, }
270+
271+
/-- Given submodules `p ⊆ M`, `p₂ ⊆ M₂`, `p₃ ⊆ M₃` and maps `f : M → M₂`, `g : M₂ → M₃` inducing
272+
`mapq f : M ⧸ p → M₂ ⧸ p₂` and `mapq g : M₂ ⧸ p₂ → M₃ ⧸ p₃` then
273+
`mapq (g ∘ f) = (mapq g) ∘ (mapq f)`. -/
274+
lemma mapq_comp {R₃ M₃ : Type*} [ring R₃] [add_comm_group M₃] [module R₃ M₃]
275+
(p₂ : submodule R₂ M₂) (p₃ : submodule R₃ M₃)
276+
{τ₂₃ : R₂ →+* R₃} {τ₁₃ : R →+* R₃} [ring_hom_comp_triple τ₁₂ τ₂₃ τ₁₃]
277+
(f : M →ₛₗ[τ₁₂] M₂) (g : M₂ →ₛₗ[τ₂₃] M₃) (hf : p ≤ p₂.comap f) (hg : p₂ ≤ p₃.comap g)
278+
(h := (hf.trans (comap_mono hg))) :
279+
p.mapq p₃ (g.comp f) h = (p₂.mapq p₃ g hg).comp (p.mapq p₂ f hf) :=
280+
by { ext, simp, }
281+
282+
@[simp] lemma mapq_id (h : p ≤ p.comap linear_map.id := by simp) :
283+
p.mapq p linear_map.id h = linear_map.id :=
284+
by { ext, simp, }
285+
286+
lemma mapq_pow {f : M →ₗ[R] M} (h : p ≤ p.comap f) (k : ℕ)
287+
(h' : p ≤ p.comap (f^k) := p.le_comap_pow_of_le_comap h k) :
288+
p.mapq p (f^k) h' = (p.mapq p f h)^k :=
289+
begin
290+
induction k with k ih,
291+
{ simp [linear_map.one_eq_id], },
292+
{ simp only [linear_map.iterate_succ, ← ih],
293+
apply p.mapq_comp, },
294+
end
295+
267296
theorem comap_liftq (f : M →ₛₗ[τ₁₂] M₂) (h) :
268297
q.comap (p.liftq f h) = (q.comap f).map (mkq p) :=
269298
le_antisymm

src/ring_theory/nilpotent.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -180,3 +180,17 @@ begin
180180
end
181181

182182
end algebra
183+
184+
namespace module.End
185+
186+
variables {M : Type v} [ring R] [add_comm_group M] [module R M]
187+
variables {f : module.End R M} {p : submodule R M} (hp : p ≤ p.comap f)
188+
189+
lemma is_nilpotent.mapq (hnp : is_nilpotent f) : is_nilpotent (p.mapq p f hp) :=
190+
begin
191+
obtain ⟨k, hk⟩ := hnp,
192+
use k,
193+
simp [← p.mapq_pow, hk],
194+
end
195+
196+
end module.End

0 commit comments

Comments
 (0)