@@ -41,6 +41,18 @@ def lower_central_series (k : ℕ) : lie_submodule R L M := (λ I, ⁅(⊤ : lie
41
41
lower_central_series R L M (k + 1 ) = ⁅(⊤ : lie_ideal R L), lower_central_series R L M k⁆ :=
42
42
function.iterate_succ_apply' (λ I, ⁅(⊤ : lie_ideal R L), I⁆) k ⊤
43
43
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
+
44
56
lemma trivial_iff_lower_central_eq_bot : is_trivial L M ↔ lower_central_series R L M 1 = ⊥ :=
45
57
begin
46
58
split; intros h,
@@ -140,6 +152,79 @@ begin
140
152
exact map_lower_central_series_le k (lie_submodule.quotient.mk' N),
141
153
end
142
154
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
+
143
228
end lie_module
144
229
145
230
@[priority 100 ]
0 commit comments