Skip to content

Commit 632ea77

Browse files
committed
chore: delete LinearMap.extendScalars which duplicates LinearMap.baseChange (#8617)
For consistency, we also rename `Submodule.extendScalars` to `Submodule.baseChange` and likewise for `LieSubmodule`.
1 parent 4b97b73 commit 632ea77

File tree

6 files changed

+60
-73
lines changed

6 files changed

+60
-73
lines changed

Mathlib/Algebra/Lie/BaseChange.lean

Lines changed: 29 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Oliver Nash
66
import Mathlib.Algebra.Algebra.RestrictScalars
77
import Mathlib.Algebra.Lie.TensorProduct
88
import Mathlib.LinearAlgebra.TensorProduct.Tower
9+
import Mathlib.RingTheory.TensorProduct
910

1011
#align_import algebra.lie.base_change from "leanprover-community/mathlib"@"9264b15ee696b7ca83f13c8ad67c83d6eb70b730"
1112

@@ -152,8 +153,8 @@ variable [CommRing R] [LieRing L] [LieAlgebra R L]
152153
[CommRing A] [Algebra R A]
153154

154155
@[simp]
155-
lemma LieModule.toEndomorphism_extendScalars (x : L) :
156-
toEndomorphism A (A ⊗[R] L) (A ⊗[R] M) (1 ⊗ₜ x) = (toEndomorphism R L M x).extendScalars A := by
156+
lemma LieModule.toEndomorphism_baseChange (x : L) :
157+
toEndomorphism A (A ⊗[R] L) (A ⊗[R] M) (1 ⊗ₜ x) = (toEndomorphism R L M x).baseChange A := by
157158
ext; simp
158159

159160
namespace LieSubmodule
@@ -164,9 +165,11 @@ open LieModule
164165

165166
variable {R L M} in
166167
/-- If `A` is an `R`-algebra, any Lie submodule of a Lie module `M` with coefficients in `R` may be
167-
pushed forward to a Lie submodule of `A ⊗ M` with coefficients in `A`. -/
168-
def extendScalars : LieSubmodule A (A ⊗[R] L) (A ⊗[R] M) :=
169-
{ (N : Submodule R M).extendScalars A with
168+
pushed forward to a Lie submodule of `A ⊗ M` with coefficients in `A`.
169+
170+
This "base change" operation is also known as "extension of scalars". -/
171+
def baseChange : LieSubmodule A (A ⊗[R] L) (A ⊗[R] M) :=
172+
{ (N : Submodule R M).baseChange A with
170173
lie_mem := by
171174
intro x m hm
172175
simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup,
@@ -176,56 +179,56 @@ def extendScalars : LieSubmodule A (A ⊗[R] L) (A ⊗[R] M) :=
176179
· change toEndomorphism A (A ⊗[R] L) (A ⊗[R] M) _ _ ∈ _
177180
simp_rw [Finsupp.total_apply, Finsupp.sum, map_sum, map_smul, toEndomorphism_apply_apply]
178181
suffices ∀ n : (N : Submodule R M).map (TensorProduct.mk R A M 1),
179-
⁅a ⊗ₜ[R] y, (n : A ⊗[R] M)⁆ ∈ (N : Submodule R M).extendScalars A by
182+
⁅a ⊗ₜ[R] y, (n : A ⊗[R] M)⁆ ∈ (N : Submodule R M).baseChange A by
180183
exact Submodule.sum_mem _ fun n _ ↦ Submodule.smul_mem _ _ (this n)
181184
rintro ⟨-, ⟨n : M, hn : n ∈ N, rfl⟩⟩
182-
exact Submodule.tmul_mem_extendScalars_of_mem _ (N.lie_mem hn)
185+
exact Submodule.tmul_mem_baseChange_of_mem _ (N.lie_mem hn)
183186
· rw [add_lie]
184-
exact ((N : Submodule R M).extendScalars A).add_mem hy hz }
187+
exact ((N : Submodule R M).baseChange A).add_mem hy hz }
185188

186189
@[simp]
187-
lemma coe_extendScalars :
188-
(N.extendScalars A : Submodule A (A ⊗[R] M)) = (N : Submodule R M).extendScalars A :=
190+
lemma coe_baseChange :
191+
(N.baseChange A : Submodule A (A ⊗[R] M)) = (N : Submodule R M).baseChange A :=
189192
rfl
190193

191194
variable {N}
192195

193196
variable {R A L M} in
194-
lemma tmul_mem_extendScalars_of_mem (a : A) {m : M} (hm : m ∈ N) :
195-
a ⊗ₜ[R] m ∈ N.extendScalars A :=
196-
(N : Submodule R M).tmul_mem_extendScalars_of_mem a hm
197+
lemma tmul_mem_baseChange_of_mem (a : A) {m : M} (hm : m ∈ N) :
198+
a ⊗ₜ[R] m ∈ N.baseChange A :=
199+
(N : Submodule R M).tmul_mem_baseChange_of_mem a hm
197200

198-
lemma mem_extendScalars_iff {m : A ⊗[R] M} :
199-
m ∈ N.extendScalars A ↔
201+
lemma mem_baseChange_iff {m : A ⊗[R] M} :
202+
m ∈ N.baseChange A ↔
200203
m ∈ Submodule.span A ((N : Submodule R M).map (TensorProduct.mk R A M 1)) :=
201204
Iff.rfl
202205

203206
@[simp]
204-
lemma extendScalars_bot : (⊥ : LieSubmodule R L M).extendScalars A = ⊥ := by
205-
simp only [extendScalars, bot_coeSubmodule, Submodule.extendScalars_bot,
207+
lemma baseChange_bot : (⊥ : LieSubmodule R L M).baseChange A = ⊥ := by
208+
simp only [baseChange, bot_coeSubmodule, Submodule.baseChange_bot,
206209
Submodule.bot_toAddSubmonoid]
207210
rfl
208211

209212
@[simp]
210-
lemma extendScalars_top : (⊤ : LieSubmodule R L M).extendScalars A = ⊤ := by
211-
simp only [extendScalars, top_coeSubmodule, Submodule.extendScalars_top,
213+
lemma baseChange_top : (⊤ : LieSubmodule R L M).baseChange A = ⊤ := by
214+
simp only [baseChange, top_coeSubmodule, Submodule.baseChange_top,
212215
Submodule.bot_toAddSubmonoid]
213216
rfl
214217

215-
lemma lie_extendScalars {I : LieIdeal R L} {N : LieSubmodule R L M} :
216-
⁅I, N⁆.extendScalars A = ⁅I.extendScalars A, N.extendScalars A⁆ := by
218+
lemma lie_baseChange {I : LieIdeal R L} {N : LieSubmodule R L M} :
219+
⁅I, N⁆.baseChange A = ⁅I.baseChange A, N.baseChange A⁆ := by
217220
set s : Set (A ⊗[R] M) := { m | ∃ x ∈ I, ∃ n ∈ N, 1 ⊗ₜ ⁅x, n⁆ = m}
218221
have : (TensorProduct.mk R A M 1) '' {m | ∃ x ∈ I, ∃ n ∈ N, ⁅x, n⁆ = m} = s := by ext; simp
219-
rw [← coe_toSubmodule_eq_iff, coe_extendScalars, lieIdeal_oper_eq_linear_span',
220-
Submodule.extendScalars_span, this, lieIdeal_oper_eq_linear_span']
222+
rw [← coe_toSubmodule_eq_iff, coe_baseChange, lieIdeal_oper_eq_linear_span',
223+
Submodule.baseChange_span, this, lieIdeal_oper_eq_linear_span']
221224
refine le_antisymm (Submodule.span_mono ?_) (Submodule.span_le.mpr ?_)
222225
· rintro - ⟨x, hx, m, hm, rfl⟩
223-
exact ⟨1 ⊗ₜ x, tmul_mem_extendScalars_of_mem 1 hx,
224-
1 ⊗ₜ m, tmul_mem_extendScalars_of_mem 1 hm, by simp⟩
226+
exact ⟨1 ⊗ₜ x, tmul_mem_baseChange_of_mem 1 hx,
227+
1 ⊗ₜ m, tmul_mem_baseChange_of_mem 1 hm, by simp⟩
225228
· rintro - ⟨x, hx, m, hm, rfl⟩
226229
revert m
227230
apply Submodule.span_induction
228-
(p := fun x' ↦ ∀ m' ∈ N.extendScalars A, ⁅x', m'⁆ ∈ Submodule.span A s) hx
231+
(p := fun x' ↦ ∀ m' ∈ N.baseChange A, ⁅x', m'⁆ ∈ Submodule.span A s) hx
229232
· rintro _ ⟨y : L, hy : y ∈ I, rfl⟩ m hm
230233
apply Submodule.span_induction (p := fun m' ↦ ⁅(1 : A) ⊗ₜ[R] y, m'⁆ ∈ Submodule.span A s) hm
231234
· rintro - ⟨m', hm' : m' ∈ N, rfl⟩

Mathlib/Algebra/Lie/Killing.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -238,12 +238,12 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm :
238238
have _i : NoZeroSMulDivisors R A := NoZeroSMulDivisors.trans R (FractionRing R) A
239239
rw [← map_zero (algebraMap R A)] at this
240240
exact NoZeroSMulDivisors.algebraMap_injective R A this
241-
rw [← LinearMap.trace_extendScalars, LinearMap.extendScalars_comp, ← toEndomorphism_extendScalars,
242-
toEndomorphism_extendScalars]
241+
rw [← LinearMap.trace_baseChange, LinearMap.baseChange_comp, ← toEndomorphism_baseChange,
242+
toEndomorphism_baseChange]
243243
replace hz : 1 ⊗ₜ z ∈ lowerCentralSeries A (A ⊗[R] L) (A ⊗[R] L) 1 := by
244244
simp only [lowerCentralSeries_succ, lowerCentralSeries_zero] at hz ⊢
245-
rw [← LieSubmodule.extendScalars_top, ← LieSubmodule.lie_extendScalars]
246-
exact Submodule.tmul_mem_extendScalars_of_mem 1 hz
245+
rw [← LieSubmodule.baseChange_top, ← LieSubmodule.lie_baseChange]
246+
exact Submodule.tmul_mem_baseChange_of_mem 1 hz
247247
replace hzc : 1 ⊗ₜ[R] z ∈ LieAlgebra.center A (A ⊗[R] L) := by
248248
simp only [mem_maxTrivSubmodule] at hzc ⊢
249249
intro y

Mathlib/Algebra/Lie/Nilpotent.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -866,11 +866,11 @@ variable (R A L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L]
866866
[CommRing A] [Algebra R A]
867867

868868
@[simp]
869-
lemma LieSubmodule.lowerCentralSeries_tensor_eq_extendScalars (k : ℕ) :
869+
lemma LieSubmodule.lowerCentralSeries_tensor_eq_baseChange (k : ℕ) :
870870
lowerCentralSeries A (A ⊗[R] L) (A ⊗[R] M) k =
871-
(lowerCentralSeries R L M k).extendScalars A := by
871+
(lowerCentralSeries R L M k).baseChange A := by
872872
induction' k with k ih; simp
873-
simp only [lowerCentralSeries_succ, ih, ← extendScalars_top, lie_extendScalars]
873+
simp only [lowerCentralSeries_succ, ih, ← baseChange_top, lie_baseChange]
874874

875875
instance LieModule.instIsNilpotentTensor [IsNilpotent R L M] :
876876
IsNilpotent A (A ⊗[R] L) (A ⊗[R] M) := by

Mathlib/LinearAlgebra/TensorProduct/Tower.lean

Lines changed: 13 additions & 35 deletions
Original file line numberDiff line numberDiff line change
@@ -451,30 +451,6 @@ end AlgebraTensorModule
451451

452452
end TensorProduct
453453

454-
namespace LinearMap
455-
456-
open TensorProduct
457-
458-
variable {R M₁ M₂ M₃ : Type*} (A : Type*) [CommSemiring R] [Semiring A] [Algebra R A]
459-
[AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃]
460-
(f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃)
461-
462-
/-- If `A` is an `R`-algebra, any `R`-linear map `M₁ → M₂` naturally induces an `A`-linear map
463-
`A ⊗ M₁ → A ⊗ M₂`. -/
464-
def extendScalars : A ⊗[R] M₁ →ₗ[A] A ⊗[R] M₂ :=
465-
TensorProduct.AlgebraTensorModule.map LinearMap.id f
466-
467-
@[simp]
468-
lemma extendScalars_tmul (a : A) (m : M₁) :
469-
f.extendScalars A (a ⊗ₜ m) = a ⊗ₜ f m :=
470-
rfl
471-
472-
lemma extendScalars_comp :
473-
(g ∘ₗ f).extendScalars A = g.extendScalars A ∘ₗ f.extendScalars A := by
474-
ext; simp
475-
476-
end LinearMap
477-
478454
namespace Submodule
479455

480456
open TensorProduct
@@ -483,32 +459,34 @@ variable {R M : Type*} (A : Type*) [CommSemiring R] [Semiring A] [Algebra R A]
483459
[AddCommMonoid M] [Module R M] (p : Submodule R M)
484460

485461
/-- If `A` is an `R`-algebra, any `R`-submodule `p` of an `R`-module `M` may be pushed forward to
486-
an `A`-submodule of `A ⊗ M`. -/
487-
def extendScalars : Submodule A (A ⊗[R] M) :=
462+
an `A`-submodule of `A ⊗ M`.
463+
464+
This "base change" operation is also known as "extension of scalars". -/
465+
def baseChange : Submodule A (A ⊗[R] M) :=
488466
span A <| p.map (TensorProduct.mk R A M 1)
489467

490468
@[simp]
491-
lemma extendScalars_bot : (⊥ : Submodule R M).extendScalars A = ⊥ := by simp [extendScalars]
469+
lemma baseChange_bot : (⊥ : Submodule R M).baseChange A = ⊥ := by simp [baseChange]
492470

493471
@[simp]
494-
lemma extendScalars_top : (⊤ : Submodule R M).extendScalars A = ⊤ := by
495-
rw [extendScalars, map_top, eq_top_iff']
472+
lemma baseChange_top : (⊤ : Submodule R M).baseChange A = ⊤ := by
473+
rw [baseChange, map_top, eq_top_iff']
496474
intro x
497475
refine x.induction_on (by simp) (fun a y ↦ ?_) (fun _ _ ↦ Submodule.add_mem _)
498476
rw [← mul_one a, ← smul_eq_mul, ← smul_tmul']
499477
refine smul_mem _ _ (subset_span ?_)
500478
simp
501479

502480
variable {A p} in
503-
lemma tmul_mem_extendScalars_of_mem (a : A) {m : M} (hm : m ∈ p) :
504-
a ⊗ₜ[R] m ∈ p.extendScalars A := by
481+
lemma tmul_mem_baseChange_of_mem (a : A) {m : M} (hm : m ∈ p) :
482+
a ⊗ₜ[R] m ∈ p.baseChange A := by
505483
rw [← mul_one a, ← smul_eq_mul, ← smul_tmul']
506-
exact smul_mem (extendScalars A p) a (subset_span ⟨m, hm, rfl⟩)
484+
exact smul_mem (baseChange A p) a (subset_span ⟨m, hm, rfl⟩)
507485

508486
@[simp]
509-
lemma extendScalars_span (s : Set M) :
510-
(span R s).extendScalars A = span A (TensorProduct.mk R A M 1 '' s) := by
511-
simp only [extendScalars, map_coe]
487+
lemma baseChange_span (s : Set M) :
488+
(span R s).baseChange A = span A (TensorProduct.mk R A M 1 '' s) := by
489+
simp only [baseChange, map_coe]
512490
refine le_antisymm (span_le.mpr ?_) (span_mono <| Set.image_subset _ subset_span)
513491
rintro - ⟨m : M, hm : m ∈ span R s, rfl⟩
514492
apply span_induction (p := fun m' ↦ (1 : A) ⊗ₜ[R] m' ∈ span A (TensorProduct.mk R A M 1 '' s)) hm

Mathlib/LinearAlgebra/Trace.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -329,9 +329,9 @@ lemma trace_comp_eq_mul_of_commute_of_isNilpotent [IsReduced R] {f g : Module.En
329329
rw [hμ, comp_add, map_add, hg, add_zero, this, LinearMap.map_smul, smul_eq_mul]
330330

331331
@[simp]
332-
lemma trace_extendScalars [Module.Free R M] [Module.Finite R M]
332+
lemma trace_baseChange [Module.Free R M] [Module.Finite R M]
333333
(f : M →ₗ[R] M) (A : Type*) [CommRing A] [Algebra R A] :
334-
trace A _ (f.extendScalars A) = algebraMap R A (trace R _ f) := by
334+
trace A _ (f.baseChange A) = algebraMap R A (trace R _ f) := by
335335
let b := Module.Free.chooseBasis R M
336336
let b' := Algebra.TensorProduct.basis A b
337337
change _ = (algebraMap R A : R →+ A) _

Mathlib/RingTheory/TensorProduct.lean

Lines changed: 9 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -64,7 +64,9 @@ variable (r : R) (f g : M →ₗ[R] N)
6464

6565
variable (A)
6666

67-
/-- `baseChange A f` for `f : M →ₗ[R] N` is the `A`-linear map `A ⊗[R] M →ₗ[A] A ⊗[R] N`. -/
67+
/-- `baseChange A f` for `f : M →ₗ[R] N` is the `A`-linear map `A ⊗[R] M →ₗ[A] A ⊗[R] N`.
68+
69+
This "base change" operation is also known as "extension of scalars". -/
6870
def baseChange (f : M →ₗ[R] N) : A ⊗[R] M →ₗ[A] A ⊗[R] N :=
6971
AlgebraTensorModule.map (LinearMap.id : A →ₗ[A] A) f
7072
#align linear_map.base_change LinearMap.baseChange
@@ -99,6 +101,10 @@ theorem baseChange_smul : (r • f).baseChange A = r • f.baseChange A := by
99101
simp [baseChange_tmul]
100102
#align linear_map.base_change_smul LinearMap.baseChange_smul
101103

104+
lemma baseChange_comp {P : Type*} [AddCommMonoid P] [Module R P] (g : N →ₗ[R] P) :
105+
(g ∘ₗ f).baseChange A = g.baseChange A ∘ₗ f.baseChange A := by
106+
ext; simp
107+
102108
variable (R A M N)
103109

104110
/-- `baseChange` as a linear map. -/
@@ -1109,8 +1115,8 @@ variable {R M₁ M₂ ι ι₂ : Type*} (A : Type*)
11091115
[AddCommGroup M₁] [Module R M₁] [AddCommGroup M₂] [Module R M₂]
11101116

11111117
@[simp]
1112-
lemma toMatrix_extendScalars (f : M₁ →ₗ[R] M₂) (b₁ : Basis ι R M₁) (b₂ : Basis ι₂ R M₂) :
1113-
toMatrix (basis A b₁) (basis A b₂) (f.extendScalars A) =
1118+
lemma toMatrix_baseChange (f : M₁ →ₗ[R] M₂) (b₁ : Basis ι R M₁) (b₂ : Basis ι₂ R M₂) :
1119+
toMatrix (basis A b₁) (basis A b₂) (f.baseChange A) =
11141120
(toMatrix b₁ b₂ f).map (algebraMap R A) := by
11151121
ext; simp [toMatrix_apply]
11161122

0 commit comments

Comments
 (0)