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

Commit 2a49f4e

Browse files
author
Oliver Nash
committed
feat(algebra/lie/direct_sum): direct sums of Lie modules (#5063)
There are three things happening here: 1. introduction of definitions of direct sums for Lie modules, 2. introduction of definitions of morphisms, equivs for Lie modules, 3. splitting out extant definition of direct sums for Lie algebras into a new file.
1 parent fee93e9 commit 2a49f4e

File tree

2 files changed

+268
-38
lines changed

2 files changed

+268
-38
lines changed

src/algebra/lie/basic.lean

Lines changed: 156 additions & 38 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ Authors: Oliver Nash
55
-/
66
import algebra.algebra.basic
77
import linear_algebra.bilinear_form
8-
import linear_algebra.direct_sum.finsupp
8+
import linear_algebra.matrix
99
import tactic.noncomm_ring
1010

1111
/-!
@@ -23,8 +23,11 @@ submodules, and the quotient of a Lie algebra by an ideal.
2323
We introduce the notation ⁅x, y⁆ for the Lie bracket. Note that these are the Unicode "square with
2424
quill" brackets rather than the usual square brackets.
2525
26-
We also introduce the notations L →ₗ⁅R⁆ L' for a morphism of Lie algebras over a commutative ring R,
27-
and L →ₗ⁅⁆ L' for the same, when the ring is implicit.
26+
Working over a fixed commutative ring `R`, we introduce the notations:
27+
* `L →ₗ⁅R⁆ L'` for a morphism of Lie algebras,
28+
* `L ≃ₗ⁅R⁆ L'` for an equivalence of Lie algebras,
29+
* `M →ₗ⁅R,L⁆ N` for a morphism of Lie algebra modules `M`, `N` over a Lie algebra `L`,
30+
* `M ≃ₗ⁅R,L⁆ N` for an equivalence of Lie algebra modules `M`, `N` over a Lie algebra `L`.
2831
2932
## Implementation notes
3033
@@ -39,7 +42,7 @@ are partially unbundled.
3942
lie bracket, ring commutator, jacobi identity, lie ring, lie algebra
4043
-/
4144

42-
universes u v w w₁
45+
universes u v w w₁ w₂
4346

4447
/-- The has_bracket class has two intended uses:
4548
@@ -65,8 +68,9 @@ identity. Forgetting the scalar multiplication, every Lie algebra is a Lie ring.
6568
@[protect_proj] class lie_algebra (R : Type u) (L : Type v) [comm_ring R] [lie_ring L] extends semimodule R L :=
6669
(lie_smul : ∀ (t : R) (x y : L), ⁅x, t • y⁆ = t • ⁅x, y⁆)
6770

68-
/-- A Lie ring module is a module over a commutative ring, together with an additive action of a
69-
Lie ring on this module, such that the Lie bracket acts as the commutator of endomorphisms. -/
71+
/-- A Lie ring module is an additive group, together with an additive action of a
72+
Lie ring on this group, such that the Lie bracket acts as the commutator of endomorphisms.
73+
(For representations of Lie *algebras* see `lie_module`.) -/
7074
@[protect_proj] class lie_ring_module (L : Type v) (M : Type w)
7175
[lie_ring L] [add_comm_group M] extends has_bracket L M :=
7276
(add_lie : ∀ (x y : L) (m : M), ⁅x + y, m⁆ = ⁅x, m⁆ + ⁅y, m⁆)
@@ -146,7 +150,6 @@ structure morphism (R : Type u) (L : Type v) (L' : Type w)
146150

147151
attribute [nolint doc_blame] lie_algebra.morphism.to_linear_map
148152

149-
infixr ` →ₗ⁅⁆ `:25 := morphism _
150153
notation L ` →ₗ⁅`:25 R:25 `⁆ `:0 L':0 := morphism R L L'
151154

152155
section morphism_properties
@@ -281,36 +284,149 @@ def trans (e₁ : L₁ ≃ₗ⁅R⁆ L₂) (e₂ : L₂ ≃ₗ⁅R⁆ L₃) : L
281284

282285
end equiv
283286

284-
namespace direct_sum
285-
open dfinsupp
286-
open_locale direct_sum
287+
end lie_algebra
287288

288-
variables {R : Type u} [comm_ring R]
289-
variables {ι : Type v} {L : ι → Type w}
290-
variables [Π i, lie_ring (L i)] [Π i, lie_algebra R (L i)]
291-
292-
/-- The direct sum of Lie rings carries a natural Lie ring structure. -/
293-
instance : lie_ring (⨁ i, L i) :=
294-
{ bracket := zip_with (λ i, λ x y, ⁅x, y⁆) (λ i, lie_zero 0),
295-
add_lie := λ x y z, by { ext, simp only [zip_with_apply, add_apply, add_lie], },
296-
lie_add := λ x y z, by { ext, simp only [zip_with_apply, add_apply, lie_add], },
297-
lie_self := λ x, by { ext, simp only [zip_with_apply, add_apply, lie_self, zero_apply], },
298-
leibniz_lie := λ x y z, by { ext, simp only [direct_sum.sub_apply,
299-
zip_with_apply, add_apply, zero_apply], apply leibniz_lie, },
300-
..(infer_instance : add_comm_group _) }
301-
302-
@[simp] lemma bracket_apply {x y : (⨁ i, L i)} {i : ι} :
303-
⁅x, y⁆ i = ⁅x i, y i⁆ := zip_with_apply _ _ x y i
304-
305-
/-- The direct sum of Lie algebras carries a natural Lie algebra structure. -/
306-
instance : lie_algebra R (⨁ i, L i) :=
307-
{ lie_smul := λ c x y, by { ext, simp only [
308-
zip_with_apply, direct_sum.smul_apply, bracket_apply, lie_smul] },
309-
..(infer_instance : module R _) }
310-
311-
end direct_sum
289+
section lie_module_morphisms
312290

313-
end lie_algebra
291+
variables (R : Type u) (L : Type v) (M : Type w) (N : Type w₁) (P : Type w₂)
292+
variables [comm_ring R] [lie_ring L] [lie_algebra R L]
293+
variables [add_comm_group M] [add_comm_group N] [add_comm_group P]
294+
variables [module R M] [module R N] [module R P]
295+
variables [lie_ring_module L M] [lie_ring_module L N] [lie_ring_module L P]
296+
variables [lie_module R L M] [lie_module R L N] [lie_module R L P]
297+
298+
set_option old_structure_cmd true
299+
300+
/-- A morphism of Lie algebra modules is a linear map which commutes with the action of the Lie
301+
algebra. -/
302+
structure lie_module_hom extends M →ₗ[R] N :=
303+
(map_lie : ∀ {x : L} {m : M}, to_fun ⁅x, m⁆ = ⁅x, to_fun m⁆)
304+
305+
attribute [nolint doc_blame] lie_module_hom.to_linear_map
306+
307+
notation M ` →ₗ⁅`:25 R,L:25 `⁆ `:0 N:0 := lie_module_hom R L M N
308+
309+
namespace lie_module_hom
310+
311+
variables {R L M N P}
312+
313+
instance : has_coe (M →ₗ⁅R,L⁆ N) (M →ₗ[R] N) := ⟨lie_module_hom.to_linear_map⟩
314+
315+
/-- see Note [function coercion] -/
316+
instance : has_coe_to_fun (M →ₗ⁅R,L⁆ N) := ⟨_, lie_module_hom.to_fun⟩
317+
318+
@[simp] lemma coe_mk (f : M → N) (h₁ h₂ h₃) :
319+
((⟨f, h₁, h₂, h₃⟩ : M →ₗ⁅R,L⁆ N) : M → N) = f := rfl
320+
321+
@[simp, norm_cast] lemma coe_to_linear_map (f : M →ₗ⁅R,L⁆ N) : ((f : M →ₗ[R] N) : M → N) = f :=
322+
rfl
323+
324+
@[simp] lemma map_lie' (f : M →ₗ⁅R,L⁆ N) (x : L) (m : M) : f ⁅x, m⁆ = ⁅x, f m⁆ :=
325+
lie_module_hom.map_lie f
326+
327+
/-- The constant 0 map is a Lie module morphism. -/
328+
instance : has_zero (M →ₗ⁅R,L⁆ N) := ⟨{ map_lie := by simp, ..(0 : M →ₗ[R] N) }⟩
329+
330+
/-- The identity map is a Lie module morphism. -/
331+
instance : has_one (M →ₗ⁅R,L⁆ M) := ⟨{ map_lie := by simp, ..(1 : M →ₗ[R] M) }⟩
332+
333+
instance : inhabited (M →ₗ⁅R,L⁆ N) := ⟨0
334+
335+
lemma coe_injective : function.injective (λ f : M →ₗ⁅R,L⁆ N, show M → N, from f) :=
336+
by { rintros ⟨f, _⟩ ⟨g, _⟩ ⟨h⟩, congr, }
337+
338+
@[ext] lemma ext {f g : M →ₗ⁅R,L⁆ N} (h : ∀ m, f m = g m) : f = g :=
339+
coe_injective $ funext h
340+
341+
lemma ext_iff {f g : M →ₗ⁅R,L⁆ N} : f = g ↔ ∀ m, f m = g m :=
342+
by { rintro rfl m, refl, }, ext⟩
343+
344+
/-- The composition of Lie module morphisms is a morphism. -/
345+
def comp (f : N →ₗ⁅R,L⁆ P) (g : M →ₗ⁅R,L⁆ N) : M →ₗ⁅R,L⁆ P :=
346+
{ map_lie := λ x m, by { change f (g ⁅x, m⁆) = ⁅x, f (g m)⁆, rw [map_lie', map_lie'], },
347+
..linear_map.comp f.to_linear_map g.to_linear_map }
348+
349+
@[simp] lemma comp_apply (f : N →ₗ⁅R,L⁆ P) (g : M →ₗ⁅R,L⁆ N) (m : M) :
350+
f.comp g m = f (g m) := rfl
351+
352+
@[norm_cast] lemma comp_coe (f : N →ₗ⁅R,L⁆ P) (g : M →ₗ⁅R,L⁆ N) :
353+
(f : N → P) ∘ (g : M → N) = f.comp g := rfl
354+
355+
/-- The inverse of a bijective morphism of Lie modules is a morphism of Lie modules. -/
356+
def inverse (f : M →ₗ⁅R,L⁆ N) (g : N → M)
357+
(h₁ : function.left_inverse g f) (h₂ : function.right_inverse g f) : N →ₗ⁅R,L⁆ M :=
358+
{ map_lie := λ x n,
359+
calc g ⁅x, n⁆ = g ⁅x, f (g n)⁆ : by rw h₂
360+
... = g (f ⁅x, g n⁆) : by rw map_lie'
361+
... = ⁅x, g n⁆ : (h₁ _),
362+
..linear_map.inverse f.to_linear_map g h₁ h₂ }
363+
364+
end lie_module_hom
365+
366+
/-- An equivalence of Lie algebra modules is a linear equivalence which is also a morphism of
367+
Lie algebra modules. -/
368+
structure lie_module_equiv extends M ≃ₗ[R] N, M →ₗ⁅R,L⁆ N
369+
370+
attribute [nolint doc_blame] lie_module_equiv.to_lie_module_hom
371+
attribute [nolint doc_blame] lie_module_equiv.to_linear_equiv
372+
373+
notation M ` ≃ₗ⁅`:25 R,L:25 `⁆ `:0 N:0 := lie_module_equiv R L M N
374+
375+
namespace lie_module_equiv
376+
377+
variables {R L M N P}
378+
379+
instance has_coe_to_lie_module_hom : has_coe (M ≃ₗ⁅R,L⁆ N) (M →ₗ⁅R,L⁆ N) := ⟨to_lie_module_hom⟩
380+
instance has_coe_to_linear_equiv : has_coe (M ≃ₗ⁅R,L⁆ N) (M ≃ₗ[R] N) := ⟨to_linear_equiv⟩
381+
382+
/-- see Note [function coercion] -/
383+
instance : has_coe_to_fun (M ≃ₗ⁅R,L⁆ N) := ⟨_, to_fun⟩
384+
385+
@[simp, norm_cast] lemma coe_to_lie_module_hom (e : M ≃ₗ⁅R,L⁆ N) : ((e : M →ₗ⁅R,L⁆ N) : M → N) = e :=
386+
rfl
387+
388+
@[simp, norm_cast] lemma coe_to_linear_equiv (e : M ≃ₗ⁅R,L⁆ N) : ((e : M ≃ₗ[R] N) : M → N) = e :=
389+
rfl
390+
391+
instance : has_one (M ≃ₗ⁅R,L⁆ M) := ⟨{ map_lie := λ x m, rfl, ..(1 : M ≃ₗ[R] M) }⟩
392+
393+
@[simp] lemma one_apply (m : M) : (1 : (M ≃ₗ⁅R,L⁆ M)) m = m := rfl
394+
395+
instance : inhabited (M ≃ₗ⁅R,L⁆ M) := ⟨1
396+
397+
/-- Lie module equivalences are reflexive. -/
398+
@[refl] def refl : M ≃ₗ⁅R,L⁆ M := 1
399+
400+
@[simp] lemma refl_apply (m : M) : (refl : M ≃ₗ⁅R,L⁆ M) m = m := rfl
401+
402+
/-- Lie module equivalences are syemmtric. -/
403+
@[symm] def symm (e : M ≃ₗ⁅R,L⁆ N) : N ≃ₗ⁅R,L⁆ M :=
404+
{ ..lie_module_hom.inverse e.to_lie_module_hom e.inv_fun e.left_inv e.right_inv,
405+
..(e : M ≃ₗ[R] N).symm }
406+
407+
@[simp] lemma symm_symm (e : M ≃ₗ⁅R,L⁆ N) : e.symm.symm = e :=
408+
by { cases e, refl, }
409+
410+
@[simp] lemma apply_symm_apply (e : M ≃ₗ⁅R,L⁆ N) : ∀ x, e (e.symm x) = x :=
411+
e.to_linear_equiv.apply_symm_apply
412+
413+
@[simp] lemma symm_apply_apply (e : M ≃ₗ⁅R,L⁆ N) : ∀ x, e.symm (e x) = x :=
414+
e.to_linear_equiv.symm_apply_apply
415+
416+
/-- Lie module equivalences are transitive. -/
417+
@[trans] def trans (e₁ : M ≃ₗ⁅R,L⁆ N) (e₂ : N ≃ₗ⁅R,L⁆ P) : M ≃ₗ⁅R,L⁆ P :=
418+
{ ..lie_module_hom.comp e₂.to_lie_module_hom e₁.to_lie_module_hom,
419+
..linear_equiv.trans e₁.to_linear_equiv e₂.to_linear_equiv }
420+
421+
@[simp] lemma trans_apply (e₁ : M ≃ₗ⁅R,L⁆ N) (e₂ : N ≃ₗ⁅R,L⁆ P) (m : M) :
422+
(e₁.trans e₂) m = e₂ (e₁ m) := rfl
423+
424+
@[simp] lemma symm_trans_apply (e₁ : M ≃ₗ⁅R,L⁆ N) (e₂ : N ≃ₗ⁅R,L⁆ P) (p : P) :
425+
(e₁.trans e₂).symm p = e₁.symm (e₂.symm p) := rfl
426+
427+
end lie_module_equiv
428+
429+
end lie_module_morphisms
314430

315431
section of_associative
316432

@@ -895,8 +1011,9 @@ begin
8951011
simp only [mem_skew_adjoint_matrices_submodule] at *,
8961012
change ⁅A, B⁆ᵀ ⬝ J = J ⬝ -⁅A, B⁆, change Aᵀ ⬝ J = J ⬝ -A at hA, change Bᵀ ⬝ J = J ⬝ -B at hB,
8971013
simp only [←matrix.mul_eq_mul] at *,
898-
rw [matrix.lie_transpose, lie_ring.of_associative_ring_bracket, lie_ring.of_associative_ring_bracket,
899-
sub_mul, mul_assoc, mul_assoc, hA, hB, ←mul_assoc, ←mul_assoc, hA, hB],
1014+
rw [matrix.lie_transpose, lie_ring.of_associative_ring_bracket,
1015+
lie_ring.of_associative_ring_bracket, sub_mul, mul_assoc, mul_assoc, hA, hB, ←mul_assoc,
1016+
←mul_assoc, hA, hB],
9001017
noncomm_ring,
9011018
end
9021019

@@ -948,7 +1065,8 @@ end
9481065
rfl
9491066

9501067
lemma mem_skew_adjoint_matrices_lie_subalgebra_unit_smul (u : units R) (J A : matrix n n R) :
951-
A ∈ skew_adjoint_matrices_lie_subalgebra ((u : R) • J) ↔ A ∈ skew_adjoint_matrices_lie_subalgebra J :=
1068+
A ∈ skew_adjoint_matrices_lie_subalgebra ((u : R) • J) ↔
1069+
A ∈ skew_adjoint_matrices_lie_subalgebra J :=
9521070
begin
9531071
change A ∈ skew_adjoint_matrices_submodule ((u : R) • J) ↔ A ∈ skew_adjoint_matrices_submodule J,
9541072
simp only [mem_skew_adjoint_matrices_submodule, matrix.is_skew_adjoint, matrix.is_adjoint_pair],

src/algebra/lie/direct_sum.lean

Lines changed: 112 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,112 @@
1+
/-
2+
Copyright (c) 2020 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 algebra.lie.basic
7+
import linear_algebra.direct_sum.finsupp
8+
9+
/-!
10+
# Direct sums of Lie algebras and Lie modules
11+
12+
Direct sums of Lie algebras and Lie modules carry natural algbebra and module structures.
13+
14+
## Tags
15+
16+
lie algebra, lie module, direct sum
17+
-/
18+
19+
universes u v w w₁
20+
21+
namespace direct_sum
22+
open dfinsupp
23+
open_locale direct_sum
24+
25+
variables {R : Type u} {ι : Type v} [comm_ring R]
26+
27+
section modules
28+
29+
/-! The direct sum of Lie modules over a fixed Lie algebra carries a natural Lie module
30+
structure. -/
31+
32+
variables {L : Type w₁} {M : ι → Type w}
33+
variables [lie_ring L] [lie_algebra R L]
34+
variables [Π i, add_comm_group (M i)] [Π i, module R (M i)]
35+
variables [Π i, lie_ring_module L (M i)] [Π i, lie_module R L (M i)]
36+
37+
instance : lie_ring_module L (⨁ i, M i) :=
38+
{ bracket := λ x m, m.map_range (λ i m', ⁅x, m'⁆) (λ i, lie_zero x),
39+
add_lie := λ x y m, by { ext, simp only [map_range_apply, add_apply, add_lie], },
40+
lie_add := λ x m n, by { ext, simp only [map_range_apply, add_apply, lie_add], },
41+
leibniz_lie := λ x y m, by { ext, simp only [map_range_apply, lie_lie, add_apply,
42+
sub_add_cancel], }, }
43+
44+
@[simp] lemma lie_module_bracket_apply (x : L) (m : ⨁ i, M i) (i : ι) :
45+
⁅x, m⁆ i = ⁅x, m i⁆ := map_range_apply _ _ m i
46+
47+
instance : lie_module R L (⨁ i, M i) :=
48+
{ smul_lie := λ t x m, by { ext i, simp only [smul_lie, lie_module_bracket_apply, smul_apply], },
49+
lie_smul := λ t x m, by { ext i, simp only [lie_smul, lie_module_bracket_apply, smul_apply], }, }
50+
51+
variables (R ι L M)
52+
53+
/-- The inclusion of each component into a direct sum as a morphism of Lie modules. -/
54+
def lie_module_of [decidable_eq ι] (j : ι) : M j →ₗ⁅R,L⁆ ⨁ i, M i :=
55+
{ map_lie := λ x m,
56+
begin
57+
ext i, by_cases h : j = i,
58+
{ rw ← h, simp, },
59+
{ simp [lof, single_eq_of_ne h], },
60+
end,
61+
..lof R ι M j }
62+
63+
/-- The projection map onto one component, as a morphism of Lie modules. -/
64+
def lie_module_component (j : ι) : (⨁ i, M i) →ₗ⁅R,L⁆ M j :=
65+
{ map_lie := λ x m,
66+
by simp only [component, lapply_apply, lie_module_bracket_apply, linear_map.to_fun_eq_coe],
67+
..component R ι M j }
68+
69+
end modules
70+
71+
section algebras
72+
73+
/-! The direct sum of Lie algebras carries a natural Lie algebra structure. -/
74+
75+
variables {L : ι → Type w}
76+
variables [Π i, lie_ring (L i)] [Π i, lie_algebra R (L i)]
77+
78+
instance : lie_ring (⨁ i, L i) :=
79+
{ bracket := zip_with (λ i, λ x y, ⁅x, y⁆) (λ i, lie_zero 0),
80+
add_lie := λ x y z, by { ext, simp only [zip_with_apply, add_apply, add_lie], },
81+
lie_add := λ x y z, by { ext, simp only [zip_with_apply, add_apply, lie_add], },
82+
lie_self := λ x, by { ext, simp only [zip_with_apply, add_apply, lie_self, zero_apply], },
83+
leibniz_lie := λ x y z, by { ext, simp only [sub_apply,
84+
zip_with_apply, add_apply, zero_apply], apply leibniz_lie, },
85+
..(infer_instance : add_comm_group _) }
86+
87+
@[simp] lemma bracket_apply (x y : ⨁ i, L i) (i : ι) :
88+
⁅x, y⁆ i = ⁅x i, y i⁆ := zip_with_apply _ _ x y i
89+
90+
instance : lie_algebra R (⨁ i, L i) :=
91+
{ lie_smul := λ c x y, by { ext, simp only [
92+
zip_with_apply, smul_apply, bracket_apply, lie_smul] },
93+
..(infer_instance : module R _) }
94+
95+
variables (R ι L)
96+
97+
/-- The inclusion of each component into the direct sum as morphism of Lie algebras. -/
98+
def lie_algebra_of [decidable_eq ι] (j : ι) : L j →ₗ⁅R⁆ ⨁ i, L i :=
99+
{ map_lie := λ x y, by
100+
{ ext i, by_cases h : j = i,
101+
{ rw ← h, simp, },
102+
{ simp [lof, single_eq_of_ne h], }, },
103+
..lof R ι L j, }
104+
105+
/-- The projection map onto one component, as a morphism of Lie algebras. -/
106+
def lie_algebra_component (j : ι) : (⨁ i, L i) →ₗ⁅R⁆ L j :=
107+
{ map_lie := λ x y, by simp only [component, bracket_apply, lapply_apply, linear_map.to_fun_eq_coe],
108+
..component R ι L j }
109+
110+
end algebras
111+
112+
end direct_sum

0 commit comments

Comments
 (0)