@@ -5,7 +5,7 @@ Authors: Oliver Nash
5
5
-/
6
6
import algebra.algebra.basic
7
7
import linear_algebra.bilinear_form
8
- import linear_algebra.direct_sum.finsupp
8
+ import linear_algebra.matrix
9
9
import tactic.noncomm_ring
10
10
11
11
/-!
@@ -23,8 +23,11 @@ submodules, and the quotient of a Lie algebra by an ideal.
23
23
We introduce the notation ⁅x, y⁆ for the Lie bracket. Note that these are the Unicode "square with
24
24
quill" brackets rather than the usual square brackets.
25
25
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`.
28
31
29
32
## Implementation notes
30
33
@@ -39,7 +42,7 @@ are partially unbundled.
39
42
lie bracket, ring commutator, jacobi identity, lie ring, lie algebra
40
43
-/
41
44
42
- universes u v w w₁
45
+ universes u v w w₁ w₂
43
46
44
47
/-- The has_bracket class has two intended uses:
45
48
@@ -65,8 +68,9 @@ identity. Forgetting the scalar multiplication, every Lie algebra is a Lie ring.
65
68
@[protect_proj] class lie_algebra (R : Type u) (L : Type v) [comm_ring R] [lie_ring L] extends semimodule R L :=
66
69
(lie_smul : ∀ (t : R) (x y : L), ⁅x, t • y⁆ = t • ⁅x, y⁆)
67
70
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`.) -/
70
74
@[protect_proj] class lie_ring_module (L : Type v) (M : Type w)
71
75
[lie_ring L] [add_comm_group M] extends has_bracket L M :=
72
76
(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)
146
150
147
151
attribute [nolint doc_blame] lie_algebra.morphism.to_linear_map
148
152
149
- infixr ` →ₗ⁅⁆ `:25 := morphism _
150
153
notation L ` →ₗ⁅`:25 R:25 `⁆ `:0 L':0 := morphism R L L'
151
154
152
155
section morphism_properties
@@ -281,36 +284,149 @@ def trans (e₁ : L₁ ≃ₗ⁅R⁆ L₂) (e₂ : L₂ ≃ₗ⁅R⁆ L₃) : L
281
284
282
285
end equiv
283
286
284
- namespace direct_sum
285
- open dfinsupp
286
- open_locale direct_sum
287
+ end lie_algebra
287
288
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
312
290
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
314
430
315
431
section of_associative
316
432
@@ -895,8 +1011,9 @@ begin
895
1011
simp only [mem_skew_adjoint_matrices_submodule] at *,
896
1012
change ⁅A, B⁆ᵀ ⬝ J = J ⬝ -⁅A, B⁆, change Aᵀ ⬝ J = J ⬝ -A at hA, change Bᵀ ⬝ J = J ⬝ -B at hB,
897
1013
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],
900
1017
noncomm_ring,
901
1018
end
902
1019
948
1065
rfl
949
1066
950
1067
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 :=
952
1070
begin
953
1071
change A ∈ skew_adjoint_matrices_submodule ((u : R) • J) ↔ A ∈ skew_adjoint_matrices_submodule J,
954
1072
simp only [mem_skew_adjoint_matrices_submodule, matrix.is_skew_adjoint, matrix.is_adjoint_pair],
0 commit comments