Skip to content

Commit 8857f4e

Browse files
committed
feat: add proof of Jordan-Chevalley-Dunford decomposition (#10295)
1 parent 9ac3278 commit 8857f4e

File tree

5 files changed

+108
-1
lines changed

5 files changed

+108
-1
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2472,6 +2472,7 @@ import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition
24722472
import Mathlib.LinearAlgebra.GeneralLinearGroup
24732473
import Mathlib.LinearAlgebra.InvariantBasisNumber
24742474
import Mathlib.LinearAlgebra.Isomorphisms
2475+
import Mathlib.LinearAlgebra.JordanChevalley
24752476
import Mathlib.LinearAlgebra.Lagrange
24762477
import Mathlib.LinearAlgebra.LinearIndependent
24772478
import Mathlib.LinearAlgebra.LinearPMap

Mathlib/Algebra/Group/Units.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -674,6 +674,11 @@ theorem isUnit_of_mul_eq_one [CommMonoid M] (a b : M) (h : a * b = 1) : IsUnit a
674674
#align is_unit_of_mul_eq_one isUnit_of_mul_eq_one
675675
#align is_add_unit_of_add_eq_zero isAddUnit_of_add_eq_zero
676676

677+
@[to_additive]
678+
theorem isUnit_of_mul_eq_one_right [CommMonoid M] (a b : M) (h : a * b = 1) : IsUnit b := by
679+
rw [mul_comm] at h
680+
exact isUnit_of_mul_eq_one b a h
681+
677682
section Monoid
678683
variable [Monoid M] {a b : M}
679684

Mathlib/Data/Polynomial/AlgebraMap.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -304,6 +304,11 @@ theorem aeval_algHom_apply {F : Type*} [FunLike F A B] [AlgHomClass F R A B]
304304
rw [map_add, hp, hq, ← map_add, ← map_add]
305305
#align polynomial.aeval_alg_hom_apply Polynomial.aeval_algHom_apply
306306

307+
@[simp]
308+
lemma coe_aeval_mk_apply {S : Subalgebra R A} (h : x ∈ S) :
309+
(aeval (⟨x, h⟩ : S) p : A) = aeval x p :=
310+
(aeval_algHom_apply S.val (⟨x, h⟩ : S) p).symm
311+
307312
theorem aeval_algEquiv (f : A ≃ₐ[R] B) (x : A) : aeval (f x) = (f : A →ₐ[R] B).comp (aeval x) :=
308313
aeval_algHom (f : A →ₐ[R] B) x
309314
#align polynomial.aeval_alg_equiv Polynomial.aeval_algEquiv
@@ -365,6 +370,23 @@ theorem _root_.Algebra.adjoin_singleton_eq_range_aeval (x : A) :
365370
rw [← Algebra.map_top, ← adjoin_X, AlgHom.map_adjoin, Set.image_singleton, aeval_X]
366371
#align algebra.adjoin_singleton_eq_range_aeval Algebra.adjoin_singleton_eq_range_aeval
367372

373+
@[simp]
374+
theorem aeval_mem_adjoin_singleton :
375+
aeval x p ∈ Algebra.adjoin R {x} := by
376+
simpa only [Algebra.adjoin_singleton_eq_range_aeval] using Set.mem_range_self p
377+
378+
instance instCommSemiringAdjoinSingleton :
379+
CommSemiring <| Algebra.adjoin R {x} :=
380+
{ mul_comm := fun ⟨p, hp⟩ ⟨q, hq⟩ ↦ by
381+
obtain ⟨p', rfl⟩ := Algebra.adjoin_singleton_eq_range_aeval R x ▸ hp
382+
obtain ⟨q', rfl⟩ := Algebra.adjoin_singleton_eq_range_aeval R x ▸ hq
383+
simp only [AlgHom.toRingHom_eq_coe, RingHom.coe_coe, Submonoid.mk_mul_mk, ← AlgHom.map_mul,
384+
mul_comm p' q'] }
385+
386+
instance instCommRingAdjoinSingleton {R A : Type*} [CommRing R] [Ring A] [Algebra R A] (x : A) :
387+
CommRing <| Algebra.adjoin R {x} :=
388+
{ mul_comm := mul_comm }
389+
368390
variable {R}
369391

370392
section Semiring
Lines changed: 79 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,79 @@
1+
/-
2+
Copyright (c) 2024 Oliver Nash. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Oliver Nash
5+
-/
6+
import Mathlib.Algebra.Squarefree.UniqueFactorizationDomain
7+
import Mathlib.Dynamics.Newton
8+
import Mathlib.FieldTheory.Perfect
9+
import Mathlib.LinearAlgebra.Semisimple
10+
11+
/-!
12+
# Jordan-Chevalley-Dunford decomposition
13+
14+
Given a finite-dimensional linear endomorphism `f`, the Jordan-Chevalley-Dunford theorem provides a
15+
sufficient condition for there to exist a nilpotent endomorphism `n` and a semisimple endomorphism
16+
`s`, such that `f = n + s` and both `n` and `s` are polynomial expressions in `f`.
17+
18+
The condition is that there exists a separable polynomial `P` such that the endomorphism `P(f)` is
19+
nilpotent. This condition is always satisfied when the coefficients are a perfect field.
20+
21+
The proof given here uses Newton's method and is taken from Chambert-Loir's notes:
22+
[Algebre](http://webusers.imj-prg.fr/~antoine.chambert-loir/enseignement/2022-23/agreg/algebre.pdf)
23+
24+
## Main definitions / results:
25+
26+
* `Module.End.exists_isNilpotent_isSemisimple`: an endomorphism of a finite-dimensional vector
27+
space over a perfect field may be written as a sum of nilpotent and semisimple endomorphisms.
28+
Moreover these nilpotent and semisimple components are polynomial expressions in the original
29+
endomorphism.
30+
31+
## TODO
32+
33+
* Uniqueness of decomposition (once we prove that the sum of commuting semisimple endomorphims is
34+
semisimple, this will follow from `Module.End.eq_zero_of_isNilpotent_isSemisimple`).
35+
36+
-/
37+
38+
open Algebra Polynomial
39+
40+
namespace Module.End
41+
42+
variable {K V : Type*} [Field K] [AddCommGroup V] [Module K V] [FiniteDimensional K V] {f : End K V}
43+
44+
theorem exists_isNilpotent_isSemisimple_of_separable_of_dvd_pow {P : K[X]} {k : ℕ}
45+
(sep : P.Separable) (nil : minpoly K f ∣ P ^ k) :
46+
∃ᵉ (n ∈ adjoin K {f}) (s ∈ adjoin K {f}), IsNilpotent n ∧ IsSemisimple s ∧ f = n + s := by
47+
set ff : adjoin K {f} := ⟨f, self_mem_adjoin_singleton K f⟩
48+
set P' := derivative P
49+
have nil' : IsNilpotent (aeval ff P) := by
50+
use k
51+
obtain ⟨q, hq⟩ := nil
52+
rw [← AlgHom.map_pow, Subtype.ext_iff]
53+
simp [hq]
54+
have sep' : IsUnit (aeval ff P') := by
55+
obtain ⟨a, b, h⟩ : IsCoprime (P ^ k) P' := sep.pow_left
56+
replace h : (aeval f b) * (aeval f P') = 1 := by
57+
simpa only [map_add, map_mul, map_one, minpoly.dvd_iff.mp nil, mul_zero, zero_add]
58+
using (aeval f).congr_arg h
59+
refine isUnit_of_mul_eq_one_right (aeval ff b) _ (Subtype.ext_iff.mpr ?_)
60+
simpa [coe_aeval_mk_apply] using h
61+
obtain ⟨⟨s, mem⟩, ⟨⟨k, hk⟩, hss⟩, -⟩ := exists_unique_nilpotent_sub_and_aeval_eq_zero nil' sep'
62+
refine ⟨f - s, ?_, s, mem, ⟨k, ?_⟩, ?_, (sub_add_cancel f s).symm⟩
63+
· exact sub_mem (self_mem_adjoin_singleton K f) mem
64+
· rw [Subtype.ext_iff] at hk
65+
simpa using hk
66+
· replace hss : aeval s P = 0 := by rwa [Subtype.ext_iff, coe_aeval_mk_apply] at hss
67+
exact isSemisimple_of_squarefree_aeval_eq_zero sep.squarefree hss
68+
69+
/-- **Jordan-Chevalley-Dunford decomposition**: an endomorphism of a finite-dimensional vector space
70+
over a perfect field may be written as a sum of nilpotent and semisimple endomorphisms. Moreover
71+
these nilpotent and semisimple components are polynomial expressions in the original endomorphism.
72+
-/
73+
theorem exists_isNilpotent_isSemisimple [PerfectField K] :
74+
∃ᵉ (n ∈ adjoin K {f}) (s ∈ adjoin K {f}), IsNilpotent n ∧ IsSemisimple s ∧ f = n + s := by
75+
obtain ⟨g, k, sep, -, nil⟩ := exists_squarefree_dvd_pow_of_ne_zero (minpoly.ne_zero_of_finite K f)
76+
rw [← PerfectField.separable_iff_squarefree] at sep
77+
exact exists_isNilpotent_isSemisimple_of_separable_of_dvd_pow sep nil
78+
79+
end Module.End

docs/undergrad.yaml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -63,7 +63,7 @@ Linear algebra:
6363
invariant subspaces of an endomorphism: 'https://en.wikipedia.org/wiki/Invariant_subspace'
6464
generalized eigenspaces: 'Module.End.generalizedEigenspace'
6565
kernels lemma: 'https://fr.wikipedia.org/wiki/Lemme_des_noyaux'
66-
Jordan-Chevalley-Dunford decomposition: 'https://en.wikipedia.org/wiki/Jordan%E2%80%93Chevalley_decomposition'
66+
Jordan-Chevalley-Dunford decomposition: 'Module.End.exists_isNilpotent_isSemisimple'
6767
Jordan normal form: 'https://en.wikipedia.org/wiki/Jordan_normal_form'
6868
Linear representations:
6969
irreducible representation: 'https://en.wikipedia.org/wiki/Irreducible_representation'

0 commit comments

Comments
 (0)