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

Commit

Permalink
feat(algebra/module): sum_smul' (for semimodules) (#1752)
Browse files Browse the repository at this point in the history
* feat(algebra/module): sum_smul' (for semimodules)

* adding docstring

* use `classical` tactic

* moving ' name to the weaker theorem
  • Loading branch information
kbuzzard authored and mergify[bot] committed Dec 20, 2019
1 parent e875492 commit 883d974
Show file tree
Hide file tree
Showing 2 changed files with 9 additions and 2 deletions.
4 changes: 2 additions & 2 deletions src/algebra/big_operators.lean
Original file line number Diff line number Diff line change
Expand Up @@ -428,10 +428,10 @@ end comm_monoid

attribute [to_additive] prod_hom

lemma sum_smul [add_comm_monoid β] (s : finset α) (n : ℕ) (f : α → β) :
lemma sum_smul' [add_comm_monoid β] (s : finset α) (n : ℕ) (f : α → β) :
s.sum (λ x, add_monoid.smul n (f x)) = add_monoid.smul n (s.sum f) :=
@prod_pow _ (multiplicative β) _ _ _ _
attribute [to_additive sum_smul] prod_pow
attribute [to_additive sum_smul'] prod_pow

@[simp] lemma sum_const [add_comm_monoid β] (b : β) :
s.sum (λ a, b) = add_monoid.smul s.card b :=
Expand Down
7 changes: 7 additions & 0 deletions src/algebra/module.lean
Original file line number Diff line number Diff line change
Expand Up @@ -54,6 +54,13 @@ instance smul.is_add_monoid_hom {r : α} : is_add_monoid_hom (λ x : β, r • x
lemma semimodule.eq_zero_of_zero_eq_one (zero_eq_one : (0 : α) = 1) : x = 0 :=
by rw [←one_smul α x, ←zero_eq_one, zero_smul]

/-- R-linearity of finite sums of elements of an R-semimodule. -/
lemma finset.sum_smul {α : Type*} {R : Type*} [semiring R] {M : Type*} [add_comm_monoid M]
[semimodule R M] (s : finset α) (r : R) (f : α → M) :
finset.sum s (λ (x : α), (r • (f x))) = r • (finset.sum s f) :=
by classical; exact
finset.induction_on s (by simp) (by simp [_root_.smul_add] {contextual := tt})

end semimodule

section prio
Expand Down

0 comments on commit 883d974

Please sign in to comment.