Skip to content

Commit

Permalink
chore(data/finset|multiset|finsupp): reduce algebra/ imports (#7797)
Browse files Browse the repository at this point in the history


Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
Co-authored-by: sgouezel <sebastien.gouezel@univ-rennes1.fr>
Co-authored-by: Floris van Doorn <fpvdoorn@gmail.com>
  • Loading branch information
4 people committed Jun 4, 2021
1 parent 89928bc commit be183e2
Show file tree
Hide file tree
Showing 17 changed files with 61 additions and 51 deletions.
5 changes: 5 additions & 0 deletions src/algebra/big_operators/basic.lean
Expand Up @@ -1377,6 +1377,11 @@ namespace list
end list

namespace multiset

lemma abs_sum_le_sum_abs [linear_ordered_add_comm_group α] {s : multiset α} :
abs s.sum ≤ (s.map abs).sum :=
le_sum_of_subadditive _ abs_zero abs_add s

variables [decidable_eq α]

@[simp] lemma to_finset_sum_count_eq (s : multiset α) :
Expand Down
17 changes: 17 additions & 0 deletions src/algebra/big_operators/finsupp.lean
Expand Up @@ -5,6 +5,7 @@ Authors: Kenny Lau
-/

import algebra.big_operators.pi
import algebra.big_operators.ring
import data.finsupp

/-!
Expand All @@ -26,10 +27,26 @@ theorem finset.sum_apply' : (∑ k in s, f k) i = ∑ k in s, f k i :=
theorem finsupp.sum_apply' : g.sum k x = g.sum (λ i b, k i b x) :=
finset.sum_apply _ _ _

section
include h0 h1

open_locale classical

theorem finsupp.sum_sum_index' : (∑ x in s, f x).sum t = ∑ x in s, (f x).sum t :=
finset.induction_on s rfl $ λ a s has ih,
by simp_rw [finset.sum_insert has, finsupp.sum_add_index h0 h1, ih]

end

section
variables {R S : Type*} [semiring R] [semiring S]

lemma finsupp.sum_mul (b : S) (s : α →₀ R) {f : α → R → S} :
(s.sum f) * b = s.sum (λ a c, (f a c) * b) :=
by simp only [finsupp.sum, finset.sum_mul]

lemma finsupp.mul_sum (b : S) (s : α →₀ R) {f : α → R → S} :
b * (s.sum f) = s.sum (λ a c, b * (f a c)) :=
by simp only [finsupp.sum, finset.mul_sum]

end
1 change: 1 addition & 0 deletions src/algebra/gcd_monoid.lean
Expand Up @@ -9,6 +9,7 @@ TODO: Generalize GCD monoid to not require normalization in all cases
-/
import algebra.associated
import data.nat.gcd
import algebra.group_power.lemmas

/-!
Expand Down
16 changes: 0 additions & 16 deletions src/algebra/group_power/lemmas.lean
Expand Up @@ -35,18 +35,6 @@ add_monoid_hom.eq_nat_cast
⟨λ n, n • (1 : A), zero_nsmul _, λ _ _, add_nsmul _ _ _⟩
(one_nsmul _)

@[simp, priority 500]
theorem list.prod_repeat (a : M) (n : ℕ) : (list.repeat a n).prod = a ^ n :=
begin
induction n with n ih,
{ rw pow_zero, refl },
{ rw [list.repeat_succ, list.prod_cons, ih, pow_succ] }
end

@[simp, priority 500]
theorem list.sum_repeat : ∀ (a : A) (n : ℕ), (list.repeat a n).sum = n • a :=
@list.prod_repeat (multiplicative A) _

@[simp, norm_cast] lemma units.coe_pow (u : units M) (n : ℕ) : ((u ^ n : units M) : M) = u ^ n :=
(units.coe_hom M).map_pow u n

Expand Down Expand Up @@ -85,10 +73,6 @@ end

end monoid

theorem nat.nsmul_eq_mul (m n : ℕ) : m • n = m * n :=
by induction m with m ih; [rw [zero_nsmul, zero_mul],
rw [succ_nsmul', ih, nat.succ_mul]]

section group
variables [group G] [group H] [add_group A] [add_group B]

Expand Down
1 change: 0 additions & 1 deletion src/algebra/iterate_hom.lean
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Yury Kudryashov
-/

import algebra.group_power
import algebra.group_power.basic
import logic.function.iterate
import group_theory.perm.basic
Expand Down
1 change: 1 addition & 0 deletions src/algebra/monoid_algebra.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Yury G. Kudryashov, Scott Morrison
-/
import algebra.algebra.basic
import algebra.big_operators.finsupp
import linear_algebra.finsupp

/-!
Expand Down
14 changes: 1 addition & 13 deletions src/data/finsupp/basic.lean
Expand Up @@ -3,10 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Scott Morrison
-/
import algebra.big_operators.order
import algebra.module.pi
import group_theory.submonoid.basic
import algebra.big_operators.ring
import data.finset.preimage
import algebra.indicator_function

Expand Down Expand Up @@ -2104,15 +2100,7 @@ instance [semiring R] [add_comm_monoid M] [module R M] {ι : Type*}
(λ i, (smul_eq_zero.mp (finsupp.ext_iff.mp h i)).resolve_left hc))⟩

section
variables [semiring R] [semiring S]

lemma sum_mul (b : S) (s : α →₀ R) {f : α → R → S} :
(s.sum f) * b = s.sum (λ a c, (f a c) * b) :=
by simp only [finsupp.sum, finset.sum_mul]

lemma mul_sum (b : S) (s : α →₀ R) {f : α → R → S} :
b * (s.sum f) = s.sum (λ a c, b * (f a c)) :=
by simp only [finsupp.sum, finset.mul_sum]
variables [has_zero R]

instance unique_of_right [subsingleton R] : unique (α →₀ R) :=
{ uniq := λ l, ext $ λ i, subsingleton.elim _ _,
Expand Down
1 change: 0 additions & 1 deletion src/data/finsupp/lattice.lean
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Aaron Anderson
-/
import data.finsupp.basic
import algebra.ordered_group

/-!
# Lattice structure on finsupps
Expand Down
15 changes: 14 additions & 1 deletion src/data/list/basic.lean
Expand Up @@ -3,10 +3,10 @@ Copyright (c) 2014 Parikshit Khanna. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro
-/
import algebra.order_functions
import control.monad.basic
import data.nat.basic
import order.rel_classes
import algebra.group_power.basic

/-!
# Basic properties of lists
Expand Down Expand Up @@ -2250,6 +2250,19 @@ theorem prod_cons : (a::l).prod = a * l.prod :=
calc (a::l).prod = foldl (*) (a * 1) l : by simp only [list.prod, foldl_cons, one_mul, mul_one]
... = _ : foldl_assoc

@[simp, priority 500]
theorem prod_repeat (a : α) (n : ℕ) : (list.repeat a n).prod = a ^ n :=
begin
induction n with n ih,
{ rw pow_zero, refl },
{ rw [list.repeat_succ, list.prod_cons, ih, pow_succ] }
end

@[simp, priority 500]
theorem sum_repeat {α : Type*} [add_monoid α] :
∀ (a : α) (n : ℕ), (list.repeat a n).sum = n • a :=
@list.prod_repeat (multiplicative α) _

@[simp, to_additive]
theorem prod_append : (l₁ ++ l₂).prod = l₁.prod * l₂.prod :=
calc (l₁ ++ l₂).prod = foldl (*) (foldl (*) 1 l₁ * 1) l₂ : by simp [list.prod]
Expand Down
12 changes: 7 additions & 5 deletions src/data/list/indexes.lean
Expand Up @@ -3,12 +3,14 @@ Copyright (c) 2020 Jannis Limperg. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jannis Limperg
-/

import data.list.basic
import data.list.defs
import data.list.zip
import data.list.range
import logic.basic

/-!
# Lemmas about list.*_with_index functions.
Some specification lemmas for `list.map_with_index`, `list.mmap_with_index`, `list.foldl_with_index`
and `list.foldr_with_index`.
-/

universes u v

Expand Down
1 change: 0 additions & 1 deletion src/data/list/perm.lean
Expand Up @@ -7,7 +7,6 @@ import data.list.bag_inter
import data.list.erase_dup
import data.list.zip
import logic.relation
import data.nat.factorial

/-!
# List permutations
Expand Down
19 changes: 7 additions & 12 deletions src/data/multiset/basic.lean
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro
-/
import data.list.perm
import algebra.group_power

/-!
# Multisets
Expand Down Expand Up @@ -833,6 +832,13 @@ prod_eq_foldl _

attribute [norm_cast] coe_prod coe_sum

@[simp, to_additive] theorem prod_to_list [comm_monoid α] (s : multiset α) :
s.to_list.prod = s.prod :=
begin
conv_rhs { rw ←coe_to_list s, },
rw coe_prod,
end

@[simp, to_additive]
theorem prod_zero [comm_monoid α] : @prod α _ 0 = 1 := rfl

Expand Down Expand Up @@ -1042,10 +1048,6 @@ lemma le_prod_nonempty_of_submultiplicative [comm_monoid α] [ordered_comm_monoi
le_prod_nonempty_of_submultiplicative_on_pred f (λ i, true) (by simp [h_mul]) (by simp) s
hs_nonempty (by simp)

lemma abs_sum_le_sum_abs [linear_ordered_field α] {s : multiset α} :
abs s.sum ≤ (s.map abs).sum :=
le_sum_of_subadditive _ abs_zero abs_add s

theorem dvd_sum [comm_semiring α] {a : α} {s : multiset α} : (∀ x ∈ s, a ∣ x) → a ∣ s.sum :=
multiset.induction_on s (λ _, dvd_zero _)
(λ x s ih h, by rw sum_cons; exact dvd_add
Expand All @@ -1054,13 +1056,6 @@ multiset.induction_on s (λ _, dvd_zero _)
@[simp] theorem sum_map_singleton (s : multiset α) : (s.map (λ a, a ::ₘ 0)).sum = s :=
multiset.induction_on s (by simp) (by simp)

@[simp, to_additive] theorem prod_to_list [comm_monoid α] (s : multiset α) :
s.to_list.prod = s.prod :=
begin
conv_rhs { rw ←coe_to_list s, },
rw coe_prod,
end

/-! ### Join -/

/-- `join S`, where `S` is a multiset of multisets, is the lift of the list join
Expand Down
2 changes: 2 additions & 0 deletions src/data/multiset/range.lean
Expand Up @@ -6,6 +6,8 @@ Authors: Mario Carneiro
import data.multiset.basic
import data.list.range

/-! # `multiset.range n` gives `{0, 1, ..., n-1}` as a multiset. -/

open list nat

namespace multiset
Expand Down
2 changes: 1 addition & 1 deletion src/data/mv_polynomial/variables.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Johan Commelin, Mario Carneiro
-/

import algebra.big_operators.order
import data.mv_polynomial.monad
import data.set.disjointed

Expand Down
3 changes: 3 additions & 0 deletions src/data/nat/basic.lean
Expand Up @@ -105,6 +105,9 @@ instance nat.subtype.semilattice_sup_bot (s : set ℕ) [decidable_pred s] [h : n
..subtype.linear_order s,
..lattice_of_linear_order }

theorem nat.nsmul_eq_mul (m n : ℕ) : m • n = m * n :=
rfl

theorem nat.eq_of_mul_eq_mul_right {n m k : ℕ} (Hm : 0 < m) (H : n * m = k * m) : n = k :=
by rw [mul_comm n m, mul_comm k m] at H; exact nat.eq_of_mul_eq_mul_left Hm H

Expand Down
1 change: 1 addition & 0 deletions src/linear_algebra/linear_independent.lean
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2020 Anne Baanen. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp, Anne Baanen
-/
import algebra.big_operators.finsupp
import linear_algebra.finsupp
import linear_algebra.prod
import linear_algebra.pi
Expand Down
1 change: 1 addition & 0 deletions src/linear_algebra/multilinear.lean
Expand Up @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel
-/
import linear_algebra.basic
import algebra.algebra.basic
import algebra.big_operators.order
import data.fintype.sort

/-!
Expand Down

0 comments on commit be183e2

Please sign in to comment.