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

Commit 26e6d4c

Browse files
committed
feat(algebra/lie/{subalgebra,submodule}): grab bag of new lemmas (#6342)
I'm splitting these off from other work to simplify subsequent reviews. Cosmetic changes aside, the following summarises what I am proposing: New definitions: - `lie_subalgebra.of_le` Definition tweaks: - `lie_hom.range` [use coercion instead of `lie_hom.to_linear_map`] - `lie_ideal.map` [factor through `submodule.map` to avoid dropping all the way down to `set.image`] New lemmas: - `lie_ideal.coe_to_lie_subalgebra_to_submodule` - `lie_ideal.incl_range` - `lie_hom.ideal_range_eq_lie_span_range` - `lie_hom.is_ideal_morphism_iff` - `lie_subalgebra.mem_range` - `lie_subalgebra.mem_map` - `lie_subalgebra.mem_of_le` - `lie_subalgebra.of_le_eq_comap_incl` - `lie_subalgebra.exists_lie_ideal_coe_eq_iff` - `lie_subalgebra.exists_nested_lie_ideal_coe_eq_iff` - `submodule.exists_lie_submodule_coe_eq_iff` - `lie_submodule.coe_lie_span_submodule_eq_iff` Deleted lemma: - `lie_hom.range_bracket` New simp attributes: - `lie_subalgebra.mem_top` - `lie_submodule.mem_top`
1 parent 78eb83a commit 26e6d4c

File tree

2 files changed

+114
-35
lines changed

2 files changed

+114
-35
lines changed

src/algebra/lie/subalgebra.lean

Lines changed: 43 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -130,50 +130,56 @@ def lie_subalgebra.incl (L' : lie_subalgebra R L) : L' →ₗ⁅R⁆ L :=
130130
{ map_lie' := λ x y, by { rw [linear_map.to_fun_eq_coe, submodule.subtype_apply], refl, },
131131
..L'.to_submodule.subtype }
132132

133+
namespace lie_hom
134+
133135
/-- The range of a morphism of Lie algebras is a Lie subalgebra. -/
134-
def lie_hom.range : lie_subalgebra R L₂ :=
136+
def range : lie_subalgebra R L₂ :=
135137
{ lie_mem' := λ x y,
136138
show x ∈ f.to_linear_map.range → y ∈ f.to_linear_map.range → ⁅x, y⁆ ∈ f.to_linear_map.range,
137139
by { repeat { rw linear_map.mem_range }, rintros ⟨x', hx⟩ ⟨y', hy⟩, refine ⟨⁅x', y'⁆, _⟩,
138-
rw [←hx, ←hy], change f ⁅x', y'⁆ = ⁅f x', f y'⁆, rw lie_hom.map_lie, },
139-
..f.to_linear_map.range }
140-
141-
@[simp] lemma lie_hom.range_bracket (x y : f.range) :
142-
(↑⁅x, y⁆ : L₂) = ⁅(↑x : L₂), ↑y⁆ := rfl
140+
rw [←hx, ←hy], change f ⁅x', y'⁆ = ⁅f x', f y'⁆, rw map_lie, },
141+
..(f : L →ₗ[R] L₂).range }
143142

144-
@[simp] lemma lie_hom.range_coe : (f.range : set L₂) = set.range f :=
143+
@[simp] lemma range_coe : (f.range : set L₂) = set.range f :=
145144
linear_map.range_coe ↑f
146145

146+
@[simp] lemma mem_range (x : L₂) : x ∈ f.range ↔ ∃ (y : L), f y = x := linear_map.mem_range
147+
148+
end lie_hom
149+
147150
namespace lie_subalgebra
148151

149-
@[simp] lemma range_incl (L' : lie_subalgebra R L) : L'.incl.range = L' :=
150-
by { rw ← coe_to_submodule_eq_iff, exact (L' : submodule R L).range_subtype, }
152+
variables (K K' : lie_subalgebra R L) (K₂ : lie_subalgebra R L₂)
153+
154+
@[simp] lemma incl_range : K.incl.range = K :=
155+
by { rw ← coe_to_submodule_eq_iff, exact (K : submodule R L).range_subtype, }
151156

152157
/-- The image of a Lie subalgebra under a Lie algebra morphism is a Lie subalgebra of the
153158
codomain. -/
154-
def map (L' : lie_subalgebra R L) : lie_subalgebra R L₂ :=
159+
def map : lie_subalgebra R L₂ :=
155160
{ lie_mem' := λ x y hx hy, by {
156161
erw submodule.mem_map at hx, rcases hx with ⟨x', hx', hx⟩, rw ←hx,
157162
erw submodule.mem_map at hy, rcases hy with ⟨y', hy', hy⟩, rw ←hy,
158163
erw submodule.mem_map,
159-
exact ⟨⁅x', y'⁆, L'.lie_mem hx' hy', lie_hom.map_lie f x' y'⟩, },
160-
..((L' : submodule R L).map (f : L →ₗ[R] L₂))}
164+
exact ⟨⁅x', y'⁆, K.lie_mem hx' hy', f.map_lie x' y'⟩, },
165+
..((K : submodule R L).map (f : L →ₗ[R] L₂)) }
166+
167+
@[simp] lemma mem_map (x : L₂) : x ∈ K.map f ↔ ∃ (y : L), y ∈ K ∧ f y = x := submodule.mem_map
161168

162-
@[simp] lemma mem_map_submodule (e : L ≃ₗ⁅R⁆ L₂) (L' : lie_subalgebra R L) (x : L₂) :
163-
x ∈ L'.map (e : L →ₗ⁅R⁆ L₂) ↔ x ∈ (L' : submodule R L).map (e : L →ₗ[R] L₂) :=
169+
-- TODO Rename and state for homs instead of equivs.
170+
@[simp] lemma mem_map_submodule (e : L ≃ₗ⁅R⁆ L₂) (x : L₂) :
171+
x ∈ K.map (e : L →ₗ⁅R⁆ L₂) ↔ x ∈ (K : submodule R L).map (e : L →ₗ[R] L₂) :=
164172
iff.rfl
165173

166174
/-- The preimage of a Lie subalgebra under a Lie algebra morphism is a Lie subalgebra of the
167175
domain. -/
168-
def comap (L' : lie_subalgebra R L₂) : lie_subalgebra R L :=
176+
def comap : lie_subalgebra R L :=
169177
{ lie_mem' := λ x y hx hy, by
170-
{ suffices : ⁅f x, f y⁆ ∈ L', by { simp [this], }, exact L'.lie_mem hx hy, },
171-
..((L' : submodule R L₂).comap (f : L →ₗ[R] L₂)), }
178+
{ suffices : ⁅f x, f y⁆ ∈ K₂, by { simp [this], }, exact K₂.lie_mem hx hy, },
179+
..((K₂ : submodule R L₂).comap (f : L →ₗ[R] L₂)), }
172180

173181
section lattice_structure
174182

175-
variables (K K' : lie_subalgebra R L)
176-
177183
open set
178184

179185
instance : partial_order (lie_subalgebra R L) :=
@@ -201,7 +207,7 @@ instance : has_top (lie_subalgebra R L) :=
201207

202208
@[simp] lemma top_coe_submodule : ((⊤ : lie_subalgebra R L) : submodule R L) = ⊤ := rfl
203209

204-
lemma mem_top (x : L) : x ∈ (⊤ : lie_subalgebra R L) := mem_univ x
210+
@[simp] lemma mem_top (x : L) : x ∈ (⊤ : lie_subalgebra R L) := mem_univ x
205211

206212
instance : has_inf (lie_subalgebra R L) :=
207213
⟨λ K K', { lie_mem' := λ x y hx hy, mem_inter (K.lie_mem hx.1 hy.1) (K'.lie_mem hx.2 hy.2),
@@ -294,7 +300,7 @@ section nested_subalgebras
294300
variables (h : K ≤ K')
295301

296302
/-- Given two nested Lie subalgebras `K ⊆ K'`, the inclusion `K ↪ K'` is a morphism of Lie
297-
algebras.-/
303+
algebras. -/
298304
def hom_of_le : K →ₗ⁅R⁆ K' :=
299305
{ map_lie' := λ x y, rfl,
300306
..submodule.of_le h }
@@ -307,6 +313,21 @@ lemma hom_of_le_injective : function.injective (hom_of_le h) :=
307313
λ x y, by simp only [hom_of_le_apply, imp_self, subtype.mk_eq_mk, submodule.coe_eq_coe,
308314
subtype.val_eq_coe]
309315

316+
/-- Given two nested Lie subalgebras `K ⊆ K'`, we can view `K` as a Lie subalgebra of `K'`,
317+
regarded as Lie algebra in its own right. -/
318+
def of_le : lie_subalgebra R K' := (hom_of_le h).range
319+
320+
@[simp] lemma mem_of_le (x : K') : x ∈ of_le h ↔ (x : L) ∈ K :=
321+
begin
322+
simp only [of_le, hom_of_le_apply, lie_hom.mem_range],
323+
split,
324+
{ rintros ⟨y, rfl⟩, exact y.property, },
325+
{ intros h, use ⟨(x : L), h⟩, simp, },
326+
end
327+
328+
lemma of_le_eq_comap_incl : of_le h = K.comap K'.incl :=
329+
by { ext, rw mem_of_le, refl, }
330+
310331
end nested_subalgebras
311332

312333
lemma map_le_iff_le_comap {K : lie_subalgebra R L} {K' : lie_subalgebra R L₂} :
@@ -329,9 +350,7 @@ variables [comm_ring R] [lie_ring L₁] [lie_ring L₂] [lie_algebra R L₁] [li
329350
noncomputable def of_injective (f : L₁ →ₗ⁅R⁆ L₂) (h : function.injective f) :
330351
L₁ ≃ₗ⁅R⁆ f.range :=
331352
have h' : (f : L₁ →ₗ[R] L₂).ker = ⊥ := linear_map.ker_eq_bot_of_injective h,
332-
{ map_lie' := λ x y, by { apply set_coe.ext,
333-
simp only [linear_equiv.of_injective_apply, lie_hom.range_bracket],
334-
apply f.map_lie, },
353+
{ map_lie' := λ x y, by { apply set_coe.ext, simpa, },
335354
..(linear_equiv.of_injective ↑f h')}
336355

337356
@[simp] lemma of_injective_apply (f : L₁ →ₗ⁅R⁆ L₂) (h : function.injective f) (x : L₁) :

src/algebra/lie/submodule.lean

Lines changed: 71 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -123,8 +123,44 @@ instance : has_coe (lie_ideal R L) (lie_subalgebra R L) := ⟨λ I, lie_ideal_su
123123
@[norm_cast] lemma lie_ideal.coe_to_subalgebra (I : lie_ideal R L) :
124124
((I : lie_subalgebra R L) : set L) = I := rfl
125125

126+
@[norm_cast] lemma lie_ideal.coe_to_lie_subalgebra_to_submodule (I : lie_ideal R L) :
127+
((I : lie_subalgebra R L) : submodule R L) = I := rfl
128+
126129
end lie_ideal
127130

131+
variables {R M}
132+
133+
lemma submodule.exists_lie_submodule_coe_eq_iff (p : submodule R M) :
134+
(∃ (N : lie_submodule R L M), ↑N = p) ↔ ∀ (x : L) (m : M), m ∈ p → ⁅x, m⁆ ∈ p :=
135+
begin
136+
split,
137+
{ rintros ⟨N, rfl⟩, exact N.lie_mem, },
138+
{ intros h, use { lie_mem := h, ..p }, exact lie_submodule.coe_to_submodule_mk p _, },
139+
end
140+
141+
namespace lie_subalgebra
142+
143+
variables {L}
144+
145+
lemma exists_lie_ideal_coe_eq_iff (K : lie_subalgebra R L):
146+
(∃ (I : lie_ideal R L), ↑I = K) ↔ ∀ (x y : L), y ∈ K → ⁅x, y⁆ ∈ K :=
147+
begin
148+
simp only [← coe_to_submodule_eq_iff, lie_ideal.coe_to_lie_subalgebra_to_submodule,
149+
submodule.exists_lie_submodule_coe_eq_iff L],
150+
exact iff.rfl,
151+
end
152+
153+
lemma exists_nested_lie_ideal_coe_eq_iff {K K' : lie_subalgebra R L} (h : K ≤ K') :
154+
(∃ (I : lie_ideal R K'), ↑I = of_le h) ↔ ∀ (x y : L), x ∈ K' → y ∈ K → ⁅x, y⁆ ∈ K :=
155+
begin
156+
simp only [exists_lie_ideal_coe_eq_iff, coe_bracket, mem_of_le],
157+
split,
158+
{ intros h' x y hx hy, exact h' ⟨x, hx⟩ ⟨y, h hy⟩ hy, },
159+
{ rintros h' ⟨x, hx⟩ ⟨y, hy⟩ hy', exact h' x y hx hy', },
160+
end
161+
162+
end lie_subalgebra
163+
128164
end lie_submodule
129165

130166
namespace lie_submodule
@@ -169,7 +205,7 @@ instance : has_top (lie_submodule R L M) :=
169205

170206
@[simp] lemma top_coe_submodule : ((⊤ : lie_submodule R L M) : submodule R M) = ⊤ := rfl
171207

172-
lemma mem_top (x : M) : x ∈ (⊤ : lie_submodule R L M) := mem_univ x
208+
@[simp] lemma mem_top (x : M) : x ∈ (⊤ : lie_submodule R L M) := mem_univ x
173209

174210
instance : has_inf (lie_submodule R L M) :=
175211
⟨λ N N', { lie_mem := λ x m h, mem_inter (N.lie_mem h.1) (N'.lie_mem h.2),
@@ -336,6 +372,14 @@ by { rw lie_span_le, exact subset.trans h subset_lie_span, }
336372
lemma lie_span_eq : lie_span R L (N : set M) = N :=
337373
le_antisymm (lie_span_le.mpr rfl.subset) subset_lie_span
338374

375+
lemma coe_lie_span_submodule_eq_iff {p : submodule R M} :
376+
(lie_span R L (p : set M) : submodule R M) = p ↔ ∃ (N : lie_submodule R L M), ↑N = p :=
377+
begin
378+
rw p.exists_lie_submodule_coe_eq_iff L, split; intros h,
379+
{ intros x m hm, rw [← h, mem_coe_submodule], exact lie_mem _ (subset_lie_span hm), },
380+
{ rw [← coe_to_submodule_mk p h, coe_to_submodule, coe_to_submodule_eq_iff, lie_span_eq], },
381+
end
382+
339383
end lie_span
340384

341385
end lattice_structure
@@ -383,7 +427,7 @@ variables (f : L →ₗ⁅R⁆ L') (I : lie_ideal R L) (J : lie_ideal R L')
383427
Note that unlike `lie_submodule.map`, we must take the `lie_span` of the image. Mathematically
384428
this is because although `f` makes `L'` into a Lie module over `L`, in general the `L` submodules of
385429
`L'` are not the same as the ideals of `L'`. -/
386-
def map : lie_ideal R L' := lie_submodule.lie_span R L' (f '' I)
430+
def map : lie_ideal R L' := lie_submodule.lie_span R L' $ (I : submodule R L).map (f : L →ₗ[R] L')
387431

388432
/-- A morphism of Lie algebras `f : L → L'` pulls back Lie ideals of `L'` to Lie ideals of `L`.
389433
@@ -432,7 +476,7 @@ by { rw lie_submodule.le_def at h, apply lie_submodule.lie_span_mono (set.image_
432476
lemma map_of_image (h : f '' I = J) : I.map f = J :=
433477
begin
434478
apply le_antisymm,
435-
{ erw [lie_submodule.lie_span_le, h], },
479+
{ erw [lie_submodule.lie_span_le, submodule.map_coe, h], },
436480
{ rw [lie_submodule.le_def, ← h], exact lie_submodule.subset_lie_span, },
437481
end
438482

@@ -462,12 +506,28 @@ def ker : lie_ideal R L := lie_ideal.comap f ⊥
462506
/-- The range of a morphism of Lie algebras as an ideal in the codomain. -/
463507
def ideal_range : lie_ideal R L' := lie_ideal.map f ⊤
464508

509+
lemma ideal_range_eq_lie_span_range :
510+
f.ideal_range = lie_submodule.lie_span R L' f.range := rfl
511+
465512
/-- The condition that the image of a morphism of Lie algebras is an ideal. -/
466513
def is_ideal_morphism : Prop := (f.ideal_range : lie_subalgebra R L') = f.range
467514

468515
@[simp] lemma is_ideal_morphism_def :
469516
f.is_ideal_morphism ↔ (f.ideal_range : lie_subalgebra R L') = f.range := iff.rfl
470517

518+
lemma is_ideal_morphism_iff :
519+
f.is_ideal_morphism ↔ ∀ (x : L') (y : L), ∃ (z : L), ⁅x, f y⁆ = f z :=
520+
begin
521+
simp only [is_ideal_morphism_def, ideal_range_eq_lie_span_range,
522+
← lie_subalgebra.coe_to_submodule_eq_iff, ← f.range.coe_to_submodule,
523+
lie_ideal.coe_to_lie_subalgebra_to_submodule, lie_submodule.coe_lie_span_submodule_eq_iff,
524+
lie_subalgebra.mem_coe_submodule, mem_range, exists_imp_distrib,
525+
submodule.exists_lie_submodule_coe_eq_iff],
526+
split,
527+
{ intros h x y, obtain ⟨z, hz⟩ := h x (f y) y rfl, use z, exact hz.symm, },
528+
{ intros h x y z hz, obtain ⟨w, hw⟩ := h x z, use w, rw [← hw, hz], },
529+
end
530+
471531
lemma range_subset_ideal_range : (f.range : set L') ⊆ f.ideal_range := lie_submodule.subset_lie_span
472532

473533
lemma map_le_ideal_range : I.map f ≤ f.ideal_range := lie_ideal.map_mono le_top
@@ -540,7 +600,7 @@ begin
540600
apply lie_submodule.lie_span_mono,
541601
rintros x ⟨y, hy₁, hy₂⟩, rw ← hy₂,
542602
erw lie_submodule.mem_sup at hy₁, obtain ⟨z₁, hz₁, z₂, hz₂, hy⟩ := hy₁, rw ← hy,
543-
rw [f.map_add, f.mem_ker.mp hz₂, add_zero], exact ⟨z₁, hz₁, rfl⟩,
603+
rw [f.coe_to_linear_map, f.map_add, f.mem_ker.mp hz₂, add_zero], exact ⟨z₁, hz₁, rfl⟩,
544604
end
545605

546606
@[simp] lemma map_comap_eq (h : f.is_ideal_morphism) : map f (comap f J) = f.ideal_range ⊓ J :=
@@ -562,6 +622,8 @@ variables (f I J)
562622
of Lie algebras. -/
563623
def incl : I →ₗ⁅R⁆ L := (I : lie_subalgebra R L).incl
564624

625+
@[simp] lemma incl_range : I.incl.range = I := (I : lie_subalgebra R L).incl_range
626+
565627
@[simp] lemma incl_apply (x : I) : I.incl x = x := rfl
566628

567629
@[simp] lemma incl_coe : (I.incl : I →ₗ[R] L) = (I : submodule R L).subtype := rfl
@@ -575,18 +637,16 @@ by rw [← lie_submodule.coe_to_submodule_eq_iff, I.incl.ker_coe_submodule,
575637

576638
@[simp] lemma incl_ideal_range : I.incl.ideal_range = I :=
577639
begin
578-
apply le_antisymm,
579-
{ erw lie_submodule.lie_span_le, intros x hx,
580-
simp only [true_and, set.mem_image, incl_apply, set.mem_univ, lie_submodule.top_coe] at hx,
581-
obtain ⟨y, hy⟩ := hx, rw ← hy, exact y.property, },
582-
{ rw [lie_submodule.le_def, ← lie_ideal.coe_to_subalgebra, ← (I : lie_subalgebra R L).range_incl],
583-
exact I.incl.range_subset_ideal_range, },
640+
rw [lie_hom.ideal_range_eq_lie_span_range, ← lie_subalgebra.coe_to_submodule,
641+
← lie_submodule.coe_to_submodule_eq_iff, incl_range, coe_to_lie_subalgebra_to_submodule,
642+
lie_submodule.coe_lie_span_submodule_eq_iff],
643+
use I,
584644
end
585645

586646
lemma incl_is_ideal_morphism : I.incl.is_ideal_morphism :=
587647
begin
588648
rw [I.incl.is_ideal_morphism_def, incl_ideal_range],
589-
exact (I : lie_subalgebra R L).range_incl.symm,
649+
exact (I : lie_subalgebra R L).incl_range.symm,
590650
end
591651

592652
end lie_ideal

0 commit comments

Comments
 (0)