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

Commit 5bbc187

Browse files
committed
feat(algebra/lie/nilpotent): a non-trivial nilpotent Lie module has non-trivial maximal trivial submodule (#11515)
The main result is `lie_module.nontrivial_max_triv_of_is_nilpotent`
1 parent a0da4f1 commit 5bbc187

File tree

5 files changed

+127
-0
lines changed

5 files changed

+127
-0
lines changed

src/algebra/lie/abelian.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -123,6 +123,17 @@ begin
123123
exact hm x,
124124
end
125125

126+
lemma le_max_triv_iff_bracket_eq_bot {N : lie_submodule R L M} :
127+
N ≤ max_triv_submodule R L M ↔ ⁅(⊤ : lie_ideal R L), N⁆ = ⊥ :=
128+
begin
129+
refine ⟨λ h, _, λ h m hm, _⟩,
130+
{ rw [← le_bot_iff, ← ideal_oper_max_triv_submodule_eq_bot R L M ⊤],
131+
exact lie_submodule.mono_lie_right _ _ ⊤ h, },
132+
{ rw mem_max_triv_submodule,
133+
rw lie_submodule.lie_eq_bot_iff at h,
134+
exact λ x, h x (lie_submodule.mem_top x) m hm, },
135+
end
136+
126137
lemma trivial_iff_le_maximal_trivial (N : lie_submodule R L M) :
127138
is_trivial L N ↔ N ≤ max_triv_submodule R L M :=
128139
⟨ λ h m hm x, is_trivial.dcases_on h (λ h, subtype.ext_iff.mp (h x ⟨m, hm⟩)),

src/algebra/lie/ideal_operations.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -108,6 +108,14 @@ begin
108108
change x ∈ (⊥ : lie_ideal R L) at hx, rw mem_bot at hx, simp [hx],
109109
end
110110

111+
lemma lie_eq_bot_iff : ⁅I, N⁆ = ⊥ ↔ ∀ (x ∈ I) (m ∈ N), ⁅(x : L), m⁆ = 0 :=
112+
begin
113+
rw [lie_ideal_oper_eq_span, lie_submodule.lie_span_eq_bot_iff],
114+
refine ⟨λ h x hx m hm, h ⁅x, m⁆ ⟨⟨x, hx⟩, ⟨m, hm⟩, rfl⟩, _⟩,
115+
rintros h - ⟨⟨x, hx⟩, ⟨⟨n, hn⟩, rfl⟩⟩,
116+
exact h x hx n hn,
117+
end
118+
111119
lemma mono_lie (h₁ : I ≤ J) (h₂ : N ≤ N') : ⁅I, N⁆ ≤ ⁅J, N'⁆ :=
112120
begin
113121
intros m h,

src/algebra/lie/nilpotent.lean

Lines changed: 85 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -41,6 +41,18 @@ def lower_central_series (k : ℕ) : lie_submodule R L M := (λ I, ⁅(⊤ : lie
4141
lower_central_series R L M (k + 1) = ⁅(⊤ : lie_ideal R L), lower_central_series R L M k⁆ :=
4242
function.iterate_succ_apply' (λ I, ⁅(⊤ : lie_ideal R L), I⁆) k ⊤
4343

44+
lemma antitone_lower_central_series : antitone $ lower_central_series R L M :=
45+
begin
46+
intros l k,
47+
induction k with k ih generalizing l;
48+
intros h,
49+
{ exact (le_zero_iff.mp h).symm ▸ le_refl _, },
50+
{ rcases nat.of_le_succ h with hk | hk,
51+
{ rw lower_central_series_succ,
52+
exact (lie_submodule.mono_lie_right _ _ ⊤ (ih hk)).trans (lie_submodule.lie_le_right _ _), },
53+
{ exact hk.symm ▸ le_refl _, }, },
54+
end
55+
4456
lemma trivial_iff_lower_central_eq_bot : is_trivial L M ↔ lower_central_series R L M 1 = ⊥ :=
4557
begin
4658
split; intros h,
@@ -140,6 +152,79 @@ begin
140152
exact map_lower_central_series_le k (lie_submodule.quotient.mk' N),
141153
end
142154

155+
/-- Given a nilpotent Lie module `M` with lower central series `M = C₀ ≥ C₁ ≥ ⋯ ≥ Cₖ = ⊥`, this is
156+
the natural number `k` (the number of inclusions).
157+
158+
For a non-nilpotent module, we use the junk value 0. -/
159+
noncomputable def nilpotency_length : ℕ :=
160+
Inf { k | lower_central_series R L M k = ⊥ }
161+
162+
lemma nilpotency_length_eq_zero_iff [is_nilpotent R L M] :
163+
nilpotency_length R L M = 0 ↔ subsingleton M :=
164+
begin
165+
let s := { k | lower_central_series R L M k = ⊥ },
166+
have hs : s.nonempty,
167+
{ unfreezingI { obtain ⟨k, hk⟩ := (by apply_instance : is_nilpotent R L M), },
168+
exact ⟨k, hk⟩, },
169+
change Inf s = 0 ↔ _,
170+
rw [← lie_submodule.subsingleton_iff R L M, ← subsingleton_iff_bot_eq_top,
171+
← lower_central_series_zero, @eq_comm (lie_submodule R L M)],
172+
refine ⟨λ h, h ▸ nat.Inf_mem hs, λ h, _⟩,
173+
rw nat.Inf_eq_zero,
174+
exact or.inl h,
175+
end
176+
177+
lemma nilpotency_length_eq_succ_iff (k : ℕ) :
178+
nilpotency_length R L M = k + 1
179+
lower_central_series R L M (k + 1) = ⊥ ∧ lower_central_series R L M k ≠ ⊥ :=
180+
begin
181+
let s := { k | lower_central_series R L M k = ⊥ },
182+
change Inf s = k + 1 ↔ k + 1 ∈ s ∧ k ∉ s,
183+
have hs : ∀ k₁ k₂, k₁ ≤ k₂ → k₁ ∈ s → k₂ ∈ s,
184+
{ rintros k₁ k₂ h₁₂ (h₁ : lower_central_series R L M k₁ = ⊥),
185+
exact eq_bot_iff.mpr (h₁ ▸ antitone_lower_central_series R L M h₁₂), },
186+
exact nat.Inf_upward_closed_eq_succ_iff hs k,
187+
end
188+
189+
/-- Given a non-trivial nilpotent Lie module `M` with lower central series
190+
`M = C₀ ≥ C₁ ≥ ⋯ ≥ Cₖ = ⊥`, this is the `k-1`th term in the lower central series (the last
191+
non-trivial term).
192+
193+
For a trivial or non-nilpotent module, this is the bottom submodule, `⊥`. -/
194+
noncomputable def lower_central_series_last : lie_submodule R L M :=
195+
match nilpotency_length R L M with
196+
| 0 := ⊥
197+
| k + 1 := lower_central_series R L M k
198+
end
199+
200+
lemma lower_central_series_last_le_max_triv :
201+
lower_central_series_last R L M ≤ max_triv_submodule R L M :=
202+
begin
203+
rw lower_central_series_last,
204+
cases h : nilpotency_length R L M with k,
205+
{ exact bot_le, },
206+
{ rw le_max_triv_iff_bracket_eq_bot,
207+
rw [nilpotency_length_eq_succ_iff, lower_central_series_succ] at h,
208+
exact h.1, },
209+
end
210+
211+
lemma nontrivial_lower_central_series_last [nontrivial M] [is_nilpotent R L M] :
212+
nontrivial (lower_central_series_last R L M) :=
213+
begin
214+
rw [lie_submodule.nontrivial_iff_ne_bot, lower_central_series_last],
215+
cases h : nilpotency_length R L M,
216+
{ rw [nilpotency_length_eq_zero_iff, ← not_nontrivial_iff_subsingleton] at h,
217+
contradiction, },
218+
{ rw nilpotency_length_eq_succ_iff at h,
219+
exact h.2, },
220+
end
221+
222+
lemma nontrivial_max_triv_of_is_nilpotent [nontrivial M] [is_nilpotent R L M] :
223+
nontrivial (max_triv_submodule R L M) :=
224+
set.nontrivial_mono
225+
(lower_central_series_last_le_max_triv R L M)
226+
(nontrivial_lower_central_series_last R L M)
227+
143228
end lie_module
144229

145230
@[priority 100]

src/algebra/lie/submodule.lean

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -379,6 +379,18 @@ not_iff_not.mp (
379379

380380
instance [nontrivial M] : nontrivial (lie_submodule R L M) := (nontrivial_iff R L M).mpr ‹_›
381381

382+
lemma nontrivial_iff_ne_bot {N : lie_submodule R L M} : nontrivial N ↔ N ≠ ⊥ :=
383+
begin
384+
split;
385+
contrapose!,
386+
{ rintros rfl ⟨⟨m₁, h₁ : m₁ ∈ (⊥ : lie_submodule R L M)⟩,
387+
⟨m₂, h₂ : m₂ ∈ (⊥ : lie_submodule R L M)⟩, h₁₂⟩,
388+
simpa [(lie_submodule.mem_bot _).mp h₁, (lie_submodule.mem_bot _).mp h₂] using h₁₂, },
389+
{ rw [not_nontrivial_iff_subsingleton, lie_submodule.eq_bot_iff],
390+
rintros ⟨h⟩ m hm,
391+
simpa using h ⟨m, hm⟩ ⟨_, N.zero_mem⟩, },
392+
end
393+
382394
variables {R L M}
383395

384396
section inclusion_maps
@@ -462,6 +474,9 @@ protected def gi : galois_insertion (lie_span R L : set M → lie_submodule R L
462474
@[simp] lemma span_univ : lie_span R L (set.univ : set M) = ⊤ :=
463475
eq_top_iff.2 $ set_like.le_def.2 $ subset_lie_span
464476

477+
lemma lie_span_eq_bot_iff : lie_span R L s = ⊥ ↔ ∀ (m ∈ s), m = (0 : M) :=
478+
by rw [_root_.eq_bot_iff, lie_span_le, bot_coe, subset_singleton_iff]
479+
465480
variables {M}
466481

467482
lemma span_union (s t : set M) : lie_span R L (s ∪ t) = lie_span R L s ⊔ lie_span R L t :=

src/data/set/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -247,6 +247,14 @@ mt $ mem_of_subset_of_mem h
247247

248248
theorem not_subset : (¬ s ⊆ t) ↔ ∃a ∈ s, a ∉ t := by simp only [subset_def, not_forall]
249249

250+
theorem nontrivial_mono {α : Type*} {s t : set α} (h₁ : s ⊆ t) (h₂ : nontrivial s) :
251+
nontrivial t :=
252+
begin
253+
rw nontrivial_iff at h₂ ⊢,
254+
obtain ⟨⟨x, hx⟩, ⟨y, hy⟩, hxy⟩ := h₂,
255+
exact ⟨⟨x, h₁ hx⟩, ⟨y, h₁ hy⟩, by simpa using hxy⟩,
256+
end
257+
250258
/-! ### Definition of strict subsets `s ⊂ t` and basic properties. -/
251259

252260
instance : has_ssubset (set α) := ⟨(<)⟩

0 commit comments

Comments
 (0)