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

Commit 40acfb6

Browse files
committed
chore(algebra/module/linear_map): move map_sum lemmas (#18106)
The lemmas `linear_map.map_sum` and `linear_equiv.map_sum` exist only for dot notation convenience, since `linear_map`s are of `add_monoid_hom_class` and therefore the generic lemma `map_sum` can replace them. Since these lemmas are the only reason that finsets are imported to `algebra/module/linear_map`, I propose deleting that import and moving the lemmas to `linear_algebra/basic`. This should open several files for porting.
1 parent 247a102 commit 40acfb6

File tree

4 files changed

+11
-11
lines changed

4 files changed

+11
-11
lines changed

src/algebra/module/equiv.lean

Lines changed: 0 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,6 @@ linear equiv, linear equivalences, linear isomorphism, linear isomorphic
3434
-/
3535

3636
open function
37-
open_locale big_operators
3837

3938
universes u u' v w x y z
4039
variables {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*}
@@ -375,9 +374,6 @@ theorem map_smul (e : N₁ ≃ₗ[R₁] N₂) (c : R₁) (x : N₁) :
375374
e (c • x) = c • e x := map_smulₛₗ e c x
376375
omit module_N₁ module_N₂
377376

378-
@[simp] lemma map_sum {s : finset ι} (u : ι → M) : e (∑ i in s, u i) = ∑ i in s, e (u i) :=
379-
e.to_linear_map.map_sum
380-
381377
@[simp] theorem map_eq_zero_iff {x : M} : e x = 0 ↔ x = 0 :=
382378
e.to_add_equiv.map_eq_zero_iff
383379
theorem map_ne_zero_iff {x : M} : e x ≠ 0 ↔ x ≠ 0 :=

src/algebra/module/linear_map.lean

Lines changed: 3 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen,
55
Frédéric Dupuis, Heather Macbeth
66
-/
7-
import algebra.big_operators.basic
87
import algebra.hom.group_action
98
import algebra.module.pi
109
import algebra.star.basic
@@ -53,8 +52,10 @@ To ensure that composition works smoothly for semilinear maps, we use the typecl
5352
linear map
5453
-/
5554

55+
assert_not_exists submonoid
56+
assert_not_exists finset
57+
5658
open function
57-
open_locale big_operators
5859

5960
universes u u' v w x y z
6061
variables {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*}
@@ -365,10 +366,6 @@ end restrict_scalars
365366

366367
variable {R}
367368

368-
@[simp] lemma map_sum {ι} {t : finset ι} {g : ι → M} :
369-
f (∑ i in t, g i) = (∑ i in t, f (g i)) :=
370-
f.to_add_monoid_hom.map_sum _ _
371-
372369
theorem to_add_monoid_hom_injective :
373370
function.injective (to_add_monoid_hom : (M →ₛₗ[σ] M₃) → (M →+ M₃)) :=
374371
λ f g h, ext $ add_monoid_hom.congr_fun h

src/algebra/module/submodule/basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -247,7 +247,7 @@ lemma injective_subtype : injective p.subtype := subtype.coe_injective
247247

248248
/-- Note the `add_submonoid` version of this lemma is called `add_submonoid.coe_finset_sum`. -/
249249
@[simp] lemma coe_sum (x : ι → p) (s : finset ι) : ↑(∑ i in s, x i) = ∑ i in s, (x i : M) :=
250-
p.subtype.map_sum
250+
map_sum p.subtype _ _
251251

252252
section restrict_scalars
253253
variables (S) [semiring S] [module S M] [module R M] [has_smul S R] [is_scalar_tower S R M]

src/linear_algebra/basic.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -146,6 +146,10 @@ variables [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] [ring_hom_comp_tripl
146146
variables (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₃] M₃)
147147
include R R₂
148148

149+
@[simp] lemma map_sum {ι : Type*} {t : finset ι} {g : ι → M} :
150+
f (∑ i in t, g i) = (∑ i in t, f (g i)) :=
151+
f.to_add_monoid_hom.map_sum _ _
152+
149153
theorem comp_assoc (h : M₃ →ₛₗ[σ₃₄] M₄) :
150154
((h.comp g : M₂ →ₛₗ[σ₂₄] M₄).comp f : M →ₛₗ[σ₁₄] M₄)
151155
= h.comp (g.comp f : M →ₛₗ[σ₁₃] M₃) := rfl
@@ -1514,6 +1518,9 @@ variables {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R}
15141518
variables {re₁₂ : ring_hom_inv_pair σ₁₂ σ₂₁} {re₂₁ : ring_hom_inv_pair σ₂₁ σ₁₂}
15151519
variables (e e' : M ≃ₛₗ[σ₁₂] M₂)
15161520

1521+
@[simp] lemma map_sum {s : finset ι} (u : ι → M) : e (∑ i in s, u i) = ∑ i in s, e (u i) :=
1522+
e.to_linear_map.map_sum
1523+
15171524
lemma map_eq_comap {p : submodule R M} :
15181525
(p.map (e : M →ₛₗ[σ₁₂] M₂) : submodule R₂ M₂) = p.comap (e.symm : M₂ →ₛₗ[σ₂₁] M) :=
15191526
set_like.coe_injective $ by simp [e.image_eq_preimage]

0 commit comments

Comments
 (0)