|
| 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 |
0 commit comments