@@ -123,8 +123,44 @@ instance : has_coe (lie_ideal R L) (lie_subalgebra R L) := ⟨λ I, lie_ideal_su
123
123
@[norm_cast] lemma lie_ideal.coe_to_subalgebra (I : lie_ideal R L) :
124
124
((I : lie_subalgebra R L) : set L) = I := rfl
125
125
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
+
126
129
end lie_ideal
127
130
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
+
128
164
end lie_submodule
129
165
130
166
namespace lie_submodule
@@ -169,7 +205,7 @@ instance : has_top (lie_submodule R L M) :=
169
205
170
206
@[simp] lemma top_coe_submodule : ((⊤ : lie_submodule R L M) : submodule R M) = ⊤ := rfl
171
207
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
173
209
174
210
instance : has_inf (lie_submodule R L M) :=
175
211
⟨λ 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, }
336
372
lemma lie_span_eq : lie_span R L (N : set M) = N :=
337
373
le_antisymm (lie_span_le.mpr rfl.subset) subset_lie_span
338
374
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
+
339
383
end lie_span
340
384
341
385
end lattice_structure
@@ -383,7 +427,7 @@ variables (f : L →ₗ⁅R⁆ L') (I : lie_ideal R L) (J : lie_ideal R L')
383
427
Note that unlike `lie_submodule.map`, we must take the `lie_span` of the image. Mathematically
384
428
this is because although `f` makes `L'` into a Lie module over `L`, in general the `L` submodules of
385
429
`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' )
387
431
388
432
/-- A morphism of Lie algebras `f : L → L'` pulls back Lie ideals of `L'` to Lie ideals of `L`.
389
433
@@ -432,7 +476,7 @@ by { rw lie_submodule.le_def at h, apply lie_submodule.lie_span_mono (set.image_
432
476
lemma map_of_image (h : f '' I = J) : I.map f = J :=
433
477
begin
434
478
apply le_antisymm,
435
- { erw [lie_submodule.lie_span_le, h], },
479
+ { erw [lie_submodule.lie_span_le, submodule.map_coe, h], },
436
480
{ rw [lie_submodule.le_def, ← h], exact lie_submodule.subset_lie_span, },
437
481
end
438
482
@@ -462,12 +506,28 @@ def ker : lie_ideal R L := lie_ideal.comap f ⊥
462
506
/-- The range of a morphism of Lie algebras as an ideal in the codomain. -/
463
507
def ideal_range : lie_ideal R L' := lie_ideal.map f ⊤
464
508
509
+ lemma ideal_range_eq_lie_span_range :
510
+ f.ideal_range = lie_submodule.lie_span R L' f.range := rfl
511
+
465
512
/-- The condition that the image of a morphism of Lie algebras is an ideal. -/
466
513
def is_ideal_morphism : Prop := (f.ideal_range : lie_subalgebra R L') = f.range
467
514
468
515
@[simp] lemma is_ideal_morphism_def :
469
516
f.is_ideal_morphism ↔ (f.ideal_range : lie_subalgebra R L') = f.range := iff.rfl
470
517
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
+
471
531
lemma range_subset_ideal_range : (f.range : set L') ⊆ f.ideal_range := lie_submodule.subset_lie_span
472
532
473
533
lemma map_le_ideal_range : I.map f ≤ f.ideal_range := lie_ideal.map_mono le_top
@@ -540,7 +600,7 @@ begin
540
600
apply lie_submodule.lie_span_mono,
541
601
rintros x ⟨y, hy₁, hy₂⟩, rw ← hy₂,
542
602
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⟩,
544
604
end
545
605
546
606
@[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)
562
622
of Lie algebras. -/
563
623
def incl : I →ₗ⁅R⁆ L := (I : lie_subalgebra R L).incl
564
624
625
+ @[simp] lemma incl_range : I.incl.range = I := (I : lie_subalgebra R L).incl_range
626
+
565
627
@[simp] lemma incl_apply (x : I) : I.incl x = x := rfl
566
628
567
629
@[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,
575
637
576
638
@[simp] lemma incl_ideal_range : I.incl.ideal_range = I :=
577
639
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,
584
644
end
585
645
586
646
lemma incl_is_ideal_morphism : I.incl.is_ideal_morphism :=
587
647
begin
588
648
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,
590
650
end
591
651
592
652
end lie_ideal
0 commit comments