Module docstring
{"# Big operators
In this file we prove theorems about products and sums indexed by a Finset.
"}
{"# Big operators
In this file we prove theorems about products and sums indexed by a Finset.
"}
Finset.prod_eq_fold
theorem
(s : Finset ι) (f : ι → M) : ∏ i ∈ s, f i = s.fold (β := M) (· * ·) 1 f
@[to_additive]
lemma prod_eq_fold (s : Finset ι) (f : ι → M) : ∏ i ∈ s, f i = s.fold (β := M) (· * ·) 1 f := rfl
Finset.prod_cons
theorem
(h : a ∉ s) : ∏ x ∈ cons a s h, f x = f a * ∏ x ∈ s, f x
@[to_additive (attr := simp)]
theorem prod_cons (h : a ∉ s) : ∏ x ∈ cons a s h, f x = f a * ∏ x ∈ s, f x :=
fold_cons h
Finset.prod_insert
theorem
[DecidableEq ι] : a ∉ s → ∏ x ∈ insert a s, f x = f a * ∏ x ∈ s, f x
@[to_additive (attr := simp)]
theorem prod_insert [DecidableEq ι] : a ∉ s → ∏ x ∈ insert a s, f x = f a * ∏ x ∈ s, f x :=
fold_insert
Finset.prod_insert_of_eq_one_if_not_mem
theorem
[DecidableEq ι] (h : a ∉ s → f a = 1) : ∏ x ∈ insert a s, f x = ∏ x ∈ s, f x
/-- The product of `f` over `insert a s` is the same as
the product over `s`, as long as `a` is in `s` or `f a = 1`. -/
@[to_additive (attr := simp) "The sum of `f` over `insert a s` is the same as
the sum over `s`, as long as `a` is in `s` or `f a = 0`."]
theorem prod_insert_of_eq_one_if_not_mem [DecidableEq ι] (h : a ∉ s → f a = 1) :
∏ x ∈ insert a s, f x = ∏ x ∈ s, f x := by
by_cases hm : a ∈ s
· simp_rw [insert_eq_of_mem hm]
· rw [prod_insert hm, h hm, one_mul]
Finset.prod_insert_one
theorem
[DecidableEq ι] (h : f a = 1) : ∏ x ∈ insert a s, f x = ∏ x ∈ s, f x
/-- The product of `f` over `insert a s` is the same as
the product over `s`, as long as `f a = 1`. -/
@[to_additive (attr := simp) "The sum of `f` over `insert a s` is the same as
the sum over `s`, as long as `f a = 0`."]
theorem prod_insert_one [DecidableEq ι] (h : f a = 1) : ∏ x ∈ insert a s, f x = ∏ x ∈ s, f x :=
prod_insert_of_eq_one_if_not_mem fun _ => h
Finset.prod_singleton
theorem
(f : ι → M) (a : ι) : ∏ x ∈ singleton a, f x = f a
@[to_additive (attr := simp)]
theorem prod_singleton (f : ι → M) (a : ι) : ∏ x ∈ singleton a, f x = f a :=
Eq.trans fold_singleton <| mul_one _
Finset.prod_pair
theorem
[DecidableEq ι] {a b : ι} (h : a ≠ b) : (∏ x ∈ ({ a, b } : Finset ι), f x) = f a * f b
@[to_additive]
theorem prod_pair [DecidableEq ι] {a b : ι} (h : a ≠ b) :
(∏ x ∈ ({a, b} : Finset ι), f x) = f a * f b := by
rw [prod_insert (not_mem_singleton.2 h), prod_singleton]
Finset.prod_image
theorem
[DecidableEq ι] {s : Finset κ} {g : κ → ι} :
(∀ x ∈ s, ∀ y ∈ s, g x = g y → x = y) → ∏ x ∈ s.image g, f x = ∏ x ∈ s, f (g x)
@[to_additive (attr := simp)]
theorem prod_image [DecidableEq ι] {s : Finset κ} {g : κ → ι} :
(∀ x ∈ s, ∀ y ∈ s, g x = g y → x = y) → ∏ x ∈ s.image g, f x = ∏ x ∈ s, f (g x) :=
fold_image
Finset.prod_attach
theorem
(s : Finset ι) (f : ι → M) : ∏ x ∈ s.attach, f x = ∏ x ∈ s, f x
@[to_additive]
lemma prod_attach (s : Finset ι) (f : ι → M) : ∏ x ∈ s.attach, f x = ∏ x ∈ s, f x := by
classical rw [← prod_image Subtype.coe_injective.injOn, attach_image_val]
Finset.prod_congr
theorem
(h : s₁ = s₂) : (∀ x ∈ s₂, f x = g x) → s₁.prod f = s₂.prod g
@[to_additive (attr := congr)]
theorem prod_congr (h : s₁ = s₂) : (∀ x ∈ s₂, f x = g x) → s₁.prod f = s₂.prod g := by
rw [h]; exact fold_congr
Finset.prod_eq_one
theorem
(h : ∀ x ∈ s, f x = 1) : ∏ x ∈ s, f x = 1
@[to_additive]
theorem prod_eq_one (h : ∀ x ∈ s, f x = 1) : ∏ x ∈ s, f x = 1 := calc
∏ x ∈ s, f x = ∏ _x ∈ s, 1 := prod_congr rfl h
_ = 1 := prod_const_one
Finset.prod_eq_one_iff
theorem
[Subsingleton Mˣ] : ∏ i ∈ s, f i = 1 ↔ ∀ i ∈ s, f i = 1
/-- In a monoid whose only unit is `1`, a product is equal to `1` iff all factors are `1`. -/
@[to_additive (attr := simp)
"In an additive monoid whose only unit is `0`, a sum is equal to `0` iff all terms are `0`."]
lemma prod_eq_one_iff [Subsingleton Mˣ] : ∏ i ∈ s, f i∏ i ∈ s, f i = 1 ↔ ∀ i ∈ s, f i = 1 := by
induction' s using Finset.cons_induction with i s hi ih <;> simp [*]
Finset.prod_disjUnion
theorem
(h) : ∏ x ∈ s₁.disjUnion s₂ h, f x = (∏ x ∈ s₁, f x) * ∏ x ∈ s₂, f x
@[to_additive]
theorem prod_disjUnion (h) :
∏ x ∈ s₁.disjUnion s₂ h, f x = (∏ x ∈ s₁, f x) * ∏ x ∈ s₂, f x := by
refine Eq.trans ?_ (fold_disjUnion h)
rw [one_mul]
rfl
Finset.prod_disjiUnion
theorem
(s : Finset κ) (t : κ → Finset ι) (h) : ∏ x ∈ s.disjiUnion t h, f x = ∏ i ∈ s, ∏ x ∈ t i, f x
@[to_additive]
theorem prod_disjiUnion (s : Finset κ) (t : κ → Finset ι) (h) :
∏ x ∈ s.disjiUnion t h, f x = ∏ i ∈ s, ∏ x ∈ t i, f x := by
refine Eq.trans ?_ (fold_disjiUnion h)
dsimp [Finset.prod, Multiset.prod, Multiset.fold, Finset.disjUnion, Finset.fold]
congr
exact prod_const_one.symm
Finset.prod_union_inter
theorem
[DecidableEq ι] : (∏ x ∈ s₁ ∪ s₂, f x) * ∏ x ∈ s₁ ∩ s₂, f x = (∏ x ∈ s₁, f x) * ∏ x ∈ s₂, f x
@[to_additive]
theorem prod_union_inter [DecidableEq ι] :
(∏ x ∈ s₁ ∪ s₂, f x) * ∏ x ∈ s₁ ∩ s₂, f x = (∏ x ∈ s₁, f x) * ∏ x ∈ s₂, f x :=
fold_union_inter
Finset.prod_union
theorem
[DecidableEq ι] (h : Disjoint s₁ s₂) : ∏ x ∈ s₁ ∪ s₂, f x = (∏ x ∈ s₁, f x) * ∏ x ∈ s₂, f x
@[to_additive]
theorem prod_union [DecidableEq ι] (h : Disjoint s₁ s₂) :
∏ x ∈ s₁ ∪ s₂, f x = (∏ x ∈ s₁, f x) * ∏ x ∈ s₂, f x := by
rw [← prod_union_inter, disjoint_iff_inter_eq_empty.mp h]; exact (mul_one _).symm
Finset.prod_filter_mul_prod_filter_not
theorem
(s : Finset ι) (p : ι → Prop) [DecidablePred p] [∀ x, Decidable (¬p x)] (f : ι → M) :
(∏ x ∈ s with p x, f x) * ∏ x ∈ s with ¬p x, f x = ∏ x ∈ s, f x
@[to_additive]
theorem prod_filter_mul_prod_filter_not
(s : Finset ι) (p : ι → Prop) [DecidablePred p] [∀ x, Decidable (¬p x)] (f : ι → M) :
(∏ x ∈ s with p x, f x) * ∏ x ∈ s with ¬p x, f x = ∏ x ∈ s, f x := by
have := Classical.decEq ι
rw [← prod_union (disjoint_filter_filter_neg s s p), filter_union_filter_neg_eq]
Finset.prod_filter_not_mul_prod_filter
theorem
(s : Finset ι) (p : ι → Prop) [DecidablePred p] [∀ x, Decidable (¬p x)] (f : ι → M) :
(∏ x ∈ s with ¬p x, f x) * ∏ x ∈ s with p x, f x = ∏ x ∈ s, f x
@[to_additive]
lemma prod_filter_not_mul_prod_filter (s : Finset ι) (p : ι → Prop) [DecidablePred p]
[∀ x, Decidable (¬p x)] (f : ι → M) :
(∏ x ∈ s with ¬p x, f x) * ∏ x ∈ s with p x, f x = ∏ x ∈ s, f x := by
rw [mul_comm, prod_filter_mul_prod_filter_not]
Finset.prod_filter_xor
theorem
(p q : ι → Prop) [DecidablePred p] [DecidablePred q] :
(∏ x ∈ s with (Xor' (p x) (q x)), f x) = (∏ x ∈ s with (p x ∧ ¬q x), f x) * (∏ x ∈ s with (q x ∧ ¬p x), f x)
@[to_additive]
theorem prod_filter_xor (p q : ι → Prop) [DecidablePred p] [DecidablePred q] :
(∏ x ∈ s with (Xor' (p x) (q x)), f x) =
(∏ x ∈ s with (p x ∧ ¬ q x), f x) * (∏ x ∈ s with (q x ∧ ¬ p x), f x) := by
classical rw [← prod_union (disjoint_filter_and_not_filter _ _), ← filter_or]
simp only [Xor']
IsCompl.prod_mul_prod
theorem
[Fintype ι] {s t : Finset ι} (h : IsCompl s t) (f : ι → M) : (∏ i ∈ s, f i) * ∏ i ∈ t, f i = ∏ i, f i
@[to_additive]
theorem _root_.IsCompl.prod_mul_prod [Fintype ι] {s t : Finset ι} (h : IsCompl s t) (f : ι → M) :
(∏ i ∈ s, f i) * ∏ i ∈ t, f i = ∏ i, f i :=
(Finset.prod_disjUnion h.disjoint).symm.trans <| by
classical rw [Finset.disjUnion_eq_union, ← Finset.sup_eq_union, h.sup_eq_top]; rfl
Finset.prod_mul_prod_compl
theorem
[Fintype ι] [DecidableEq ι] (s : Finset ι) (f : ι → M) : (∏ i ∈ s, f i) * ∏ i ∈ sᶜ, f i = ∏ i, f i
/-- Multiplying the products of a function over `s` and over `sᶜ` gives the whole product.
For a version expressed with subtypes, see `Fintype.prod_subtype_mul_prod_subtype`. -/
@[to_additive "Adding the sums of a function over `s` and over `sᶜ` gives the whole sum.
For a version expressed with subtypes, see `Fintype.sum_subtype_add_sum_subtype`. "]
lemma prod_mul_prod_compl [Fintype ι] [DecidableEq ι] (s : Finset ι) (f : ι → M) :
(∏ i ∈ s, f i) * ∏ i ∈ sᶜ, f i = ∏ i, f i :=
IsCompl.prod_mul_prod isCompl_compl f
Finset.prod_compl_mul_prod
theorem
[Fintype ι] [DecidableEq ι] (s : Finset ι) (f : ι → M) : (∏ i ∈ sᶜ, f i) * ∏ i ∈ s, f i = ∏ i, f i
@[to_additive]
lemma prod_compl_mul_prod [Fintype ι] [DecidableEq ι] (s : Finset ι) (f : ι → M) :
(∏ i ∈ sᶜ, f i) * ∏ i ∈ s, f i = ∏ i, f i :=
(@isCompl_compl _ s _).symm.prod_mul_prod f
Finset.prod_sdiff
theorem
[DecidableEq ι] (h : s₁ ⊆ s₂) : (∏ x ∈ s₂ \ s₁, f x) * ∏ x ∈ s₁, f x = ∏ x ∈ s₂, f x
@[to_additive]
theorem prod_sdiff [DecidableEq ι] (h : s₁ ⊆ s₂) :
(∏ x ∈ s₂ \ s₁, f x) * ∏ x ∈ s₁, f x = ∏ x ∈ s₂, f x := by
rw [← prod_union sdiff_disjoint, sdiff_union_of_subset h]
Finset.prod_subset_one_on_sdiff
theorem
[DecidableEq ι] (h : s₁ ⊆ s₂) (hg : ∀ x ∈ s₂ \ s₁, g x = 1) (hfg : ∀ x ∈ s₁, f x = g x) : ∏ i ∈ s₁, f i = ∏ i ∈ s₂, g i
@[to_additive]
theorem prod_subset_one_on_sdiff [DecidableEq ι] (h : s₁ ⊆ s₂) (hg : ∀ x ∈ s₂ \ s₁, g x = 1)
(hfg : ∀ x ∈ s₁, f x = g x) : ∏ i ∈ s₁, f i = ∏ i ∈ s₂, g i := by
rw [← prod_sdiff h, prod_eq_one hg, one_mul]
exact prod_congr rfl hfg
Finset.prod_subset
theorem
(h : s₁ ⊆ s₂) (hf : ∀ x ∈ s₂, x ∉ s₁ → f x = 1) : ∏ x ∈ s₁, f x = ∏ x ∈ s₂, f x
@[to_additive]
theorem prod_subset (h : s₁ ⊆ s₂) (hf : ∀ x ∈ s₂, x ∉ s₁ → f x = 1) :
∏ x ∈ s₁, f x = ∏ x ∈ s₂, f x :=
haveI := Classical.decEq ι
prod_subset_one_on_sdiff h (by simpa) fun _ _ => rfl
Finset.prod_disj_sum
theorem
(s : Finset ι) (t : Finset κ) (f : ι ⊕ κ → M) :
∏ x ∈ s.disjSum t, f x = (∏ x ∈ s, f (Sum.inl x)) * ∏ x ∈ t, f (Sum.inr x)
@[to_additive (attr := simp)]
theorem prod_disj_sum (s : Finset ι) (t : Finset κ) (f : ι ⊕ κ → M) :
∏ x ∈ s.disjSum t, f x = (∏ x ∈ s, f (Sum.inl x)) * ∏ x ∈ t, f (Sum.inr x) := by
rw [← map_inl_disjUnion_map_inr, prod_disjUnion, prod_map, prod_map]
rfl
Finset.prod_sum_eq_prod_toLeft_mul_prod_toRight
theorem
(s : Finset (ι ⊕ κ)) (f : ι ⊕ κ → M) : ∏ x ∈ s, f x = (∏ x ∈ s.toLeft, f (Sum.inl x)) * ∏ x ∈ s.toRight, f (Sum.inr x)
@[to_additive]
lemma prod_sum_eq_prod_toLeft_mul_prod_toRight (s : Finset (ι ⊕ κ)) (f : ι ⊕ κ → M) :
∏ x ∈ s, f x = (∏ x ∈ s.toLeft, f (Sum.inl x)) * ∏ x ∈ s.toRight, f (Sum.inr x) := by
rw [← Finset.toLeft_disjSum_toRight (u := s), Finset.prod_disj_sum, Finset.toLeft_disjSum,
Finset.toRight_disjSum]
Finset.prod_sumElim
theorem
(s : Finset ι) (t : Finset κ) (f : ι → M) (g : κ → M) :
∏ x ∈ s.disjSum t, Sum.elim f g x = (∏ x ∈ s, f x) * ∏ x ∈ t, g x
@[to_additive]
theorem prod_sumElim (s : Finset ι) (t : Finset κ) (f : ι → M) (g : κ → M) :
∏ x ∈ s.disjSum t, Sum.elim f g x = (∏ x ∈ s, f x) * ∏ x ∈ t, g x := by simp
Finset.prod_biUnion
theorem
[DecidableEq ι] {s : Finset κ} {t : κ → Finset ι} (hs : Set.PairwiseDisjoint (↑s) t) :
∏ x ∈ s.biUnion t, f x = ∏ x ∈ s, ∏ i ∈ t x, f i
@[to_additive]
theorem prod_biUnion [DecidableEq ι] {s : Finset κ} {t : κ → Finset ι}
(hs : Set.PairwiseDisjoint (↑s) t) : ∏ x ∈ s.biUnion t, f x = ∏ x ∈ s, ∏ i ∈ t x, f i := by
rw [← disjiUnion_eq_biUnion _ _ hs, prod_disjiUnion]
Finset.prod_of_injOn
theorem
(e : ι → κ) (he : Set.InjOn e s) (hest : Set.MapsTo e s t) (h' : ∀ i ∈ t, i ∉ e '' s → g i = 1)
(h : ∀ i ∈ s, f i = g (e i)) : ∏ i ∈ s, f i = ∏ j ∈ t, g j
@[to_additive]
lemma prod_of_injOn (e : ι → κ) (he : Set.InjOn e s) (hest : Set.MapsTo e s t)
(h' : ∀ i ∈ t, i ∉ e '' s → g i = 1) (h : ∀ i ∈ s, f i = g (e i)) :
∏ i ∈ s, f i = ∏ j ∈ t, g j := by
classical
exact (prod_nbij e (fun a ↦ mem_image_of_mem e) he (by simp [Set.surjOn_image]) h).trans <|
prod_subset (image_subset_iff.2 hest) <| by simpa using h'
Finset.prod_fiberwise_eq_prod_filter
theorem
(s : Finset ι) (t : Finset κ) (g : ι → κ) (f : ι → M) : ∏ j ∈ t, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s with g i ∈ t, f i
@[to_additive]
lemma prod_fiberwise_eq_prod_filter (s : Finset ι) (t : Finset κ) (g : ι → κ) (f : ι → M) :
∏ j ∈ t, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s with g i ∈ t, f i := by
rw [← prod_disjiUnion, disjiUnion_filter_eq]
Finset.prod_fiberwise_eq_prod_filter'
theorem
(s : Finset ι) (t : Finset κ) (g : ι → κ) (f : κ → M) :
∏ j ∈ t, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s with g i ∈ t, f (g i)
@[to_additive]
lemma prod_fiberwise_eq_prod_filter' (s : Finset ι) (t : Finset κ) (g : ι → κ) (f : κ → M) :
∏ j ∈ t, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s with g i ∈ t, f (g i) := by
calc
_ = ∏ j ∈ t, ∏ i ∈ s with g i = j, f (g i) :=
prod_congr rfl fun j _ ↦ prod_congr rfl fun i hi ↦ by rw [(mem_filter.1 hi).2]
_ = _ := prod_fiberwise_eq_prod_filter _ _ _ _
Finset.prod_fiberwise_of_maps_to
theorem
{g : ι → κ} (h : ∀ i ∈ s, g i ∈ t) (f : ι → M) : ∏ j ∈ t, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s, f i
@[to_additive]
lemma prod_fiberwise_of_maps_to {g : ι → κ} (h : ∀ i ∈ s, g i ∈ t) (f : ι → M) :
∏ j ∈ t, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s, f i := by
rw [← prod_disjiUnion, disjiUnion_filter_eq_of_maps_to h]
Finset.prod_fiberwise_of_maps_to'
theorem
{g : ι → κ} (h : ∀ i ∈ s, g i ∈ t) (f : κ → M) : ∏ j ∈ t, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s, f (g i)
@[to_additive]
lemma prod_fiberwise_of_maps_to' {g : ι → κ} (h : ∀ i ∈ s, g i ∈ t) (f : κ → M) :
∏ j ∈ t, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s, f (g i) := by
calc
_ = ∏ j ∈ t, ∏ i ∈ s with g i = j, f (g i) :=
prod_congr rfl fun y _ ↦ prod_congr rfl fun x hx ↦ by rw [(mem_filter.1 hx).2]
_ = _ := prod_fiberwise_of_maps_to h _
Finset.prod_fiberwise
theorem
(s : Finset ι) (g : ι → κ) (f : ι → M) : ∏ j, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s, f i
@[to_additive]
lemma prod_fiberwise (s : Finset ι) (g : ι → κ) (f : ι → M) :
∏ j, ∏ i ∈ s with g i = j, f i = ∏ i ∈ s, f i :=
prod_fiberwise_of_maps_to (fun _ _ ↦ mem_univ _) _
Finset.prod_fiberwise'
theorem
(s : Finset ι) (g : ι → κ) (f : κ → M) : ∏ j, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s, f (g i)
@[to_additive]
lemma prod_fiberwise' (s : Finset ι) (g : ι → κ) (f : κ → M) :
∏ j, ∏ i ∈ s with g i = j, f j = ∏ i ∈ s, f (g i) :=
prod_fiberwise_of_maps_to' (fun _ _ ↦ mem_univ _) _
Finset.prod_diag
theorem
[DecidableEq ι] (s : Finset ι) (f : ι × ι → M) : ∏ i ∈ s.diag, f i = ∏ i ∈ s, f (i, i)
@[to_additive (attr := simp)]
lemma prod_diag [DecidableEq ι] (s : Finset ι) (f : ι × ι → M) :
∏ i ∈ s.diag, f i = ∏ i ∈ s, f (i, i) := by
apply prod_nbij' Prod.fst (fun i ↦ (i, i)) <;> simp
Finset.prod_image'
theorem
[DecidableEq ι] {s : Finset κ} {g : κ → ι} (h : κ → M) (eq : ∀ i ∈ s, f (g i) = ∏ j ∈ s with g j = g i, h j) :
∏ a ∈ s.image g, f a = ∏ i ∈ s, h i
@[to_additive]
theorem prod_image' [DecidableEq ι] {s : Finset κ} {g : κ → ι} (h : κ → M)
(eq : ∀ i ∈ s, f (g i) = ∏ j ∈ s with g j = g i, h j) :
∏ a ∈ s.image g, f a = ∏ i ∈ s, h i :=
calc
∏ a ∈ s.image g, f a = ∏ a ∈ s.image g, ∏ j ∈ s with g j = a, h j :=
(prod_congr rfl) fun _a hx =>
let ⟨i, his, hi⟩ := mem_image.1 hx
hi ▸ eq i his
_ = ∏ i ∈ s, h i := prod_fiberwise_of_maps_to (fun _ => mem_image_of_mem g) _
Finset.prod_mul_distrib
theorem
: ∏ x ∈ s, f x * g x = (∏ x ∈ s, f x) * ∏ x ∈ s, g x
@[to_additive]
theorem prod_mul_distrib : ∏ x ∈ s, f x * g x = (∏ x ∈ s, f x) * ∏ x ∈ s, g x :=
Eq.trans (by rw [one_mul]; rfl) fold_op_distrib
Finset.prod_mul_prod_comm
theorem
(f g h i : ι → M) : (∏ a ∈ s, f a * g a) * ∏ a ∈ s, h a * i a = (∏ a ∈ s, f a * h a) * ∏ a ∈ s, g a * i a
@[to_additive]
lemma prod_mul_prod_comm (f g h i : ι → M) :
(∏ a ∈ s, f a * g a) * ∏ a ∈ s, h a * i a = (∏ a ∈ s, f a * h a) * ∏ a ∈ s, g a * i a := by
simp_rw [prod_mul_distrib, mul_mul_mul_comm]
Finset.prod_filter_of_ne
theorem
{p : ι → Prop} [DecidablePred p] (hp : ∀ x ∈ s, f x ≠ 1 → p x) : ∏ x ∈ s with p x, f x = ∏ x ∈ s, f x
@[to_additive]
theorem prod_filter_of_ne {p : ι → Prop} [DecidablePred p] (hp : ∀ x ∈ s, f x ≠ 1 → p x) :
∏ x ∈ s with p x, f x = ∏ x ∈ s, f x :=
(prod_subset (filter_subset _ _)) fun x => by
classical
rw [not_imp_comm, mem_filter]
exact fun h₁ h₂ => ⟨h₁, by simpa using hp _ h₁ h₂⟩
Finset.prod_filter_ne_one
theorem
(s : Finset ι) [∀ x, Decidable (f x ≠ 1)] : ∏ x ∈ s with f x ≠ 1, f x = ∏ x ∈ s, f x
@[to_additive]
theorem prod_filter_ne_one (s : Finset ι) [∀ x, Decidable (f x ≠ 1)] :
∏ x ∈ s with f x ≠ 1, f x = ∏ x ∈ s, f x :=
prod_filter_of_ne fun _ _ => id
Finset.prod_filter
theorem
(p : ι → Prop) [DecidablePred p] (f : ι → M) : ∏ a ∈ s with p a, f a = ∏ a ∈ s, if p a then f a else 1
@[to_additive]
theorem prod_filter (p : ι → Prop) [DecidablePred p] (f : ι → M) :
∏ a ∈ s with p a, f a = ∏ a ∈ s, if p a then f a else 1 :=
calc
∏ a ∈ s with p a, f a = ∏ a ∈ s with p a, if p a then f a else 1 :=
prod_congr rfl fun a h => by rw [if_pos]; simpa using (mem_filter.1 h).2
_ = ∏ a ∈ s, if p a then f a else 1 := by
{ refine prod_subset (filter_subset _ s) fun x hs h => ?_
rw [mem_filter, not_and] at h
exact if_neg (by simpa using h hs) }
Finset.prod_eq_single_of_mem
theorem
{s : Finset ι} {f : ι → M} (a : ι) (h : a ∈ s) (h₀ : ∀ b ∈ s, b ≠ a → f b = 1) : ∏ x ∈ s, f x = f a
@[to_additive]
theorem prod_eq_single_of_mem {s : Finset ι} {f : ι → M} (a : ι) (h : a ∈ s)
(h₀ : ∀ b ∈ s, b ≠ a → f b = 1) : ∏ x ∈ s, f x = f a := by
haveI := Classical.decEq ι
calc
∏ x ∈ s, f x = ∏ x ∈ {a}, f x := by
{ refine (prod_subset ?_ ?_).symm
· intro _ H
rwa [mem_singleton.1 H]
· simpa only [mem_singleton] }
_ = f a := prod_singleton _ _
Finset.prod_eq_single
theorem
{s : Finset ι} {f : ι → M} (a : ι) (h₀ : ∀ b ∈ s, b ≠ a → f b = 1) (h₁ : a ∉ s → f a = 1) : ∏ x ∈ s, f x = f a
@[to_additive]
theorem prod_eq_single {s : Finset ι} {f : ι → M} (a : ι) (h₀ : ∀ b ∈ s, b ≠ a → f b = 1)
(h₁ : a ∉ s → f a = 1) : ∏ x ∈ s, f x = f a :=
haveI := Classical.decEq ι
by_cases (prod_eq_single_of_mem a · h₀) fun this =>
(prod_congr rfl fun b hb => h₀ b hb <| by rintro rfl; exact this hb).trans <|
prod_const_one.trans (h₁ this).symm
Finset.prod_union_eq_left
theorem
[DecidableEq ι] (hs : ∀ a ∈ s₂, a ∉ s₁ → f a = 1) : ∏ a ∈ s₁ ∪ s₂, f a = ∏ a ∈ s₁, f a
@[to_additive]
lemma prod_union_eq_left [DecidableEq ι] (hs : ∀ a ∈ s₂, a ∉ s₁ → f a = 1) :
∏ a ∈ s₁ ∪ s₂, f a = ∏ a ∈ s₁, f a :=
Eq.symm <|
prod_subset subset_union_left fun _a ha ha' ↦ hs _ ((mem_union.1 ha).resolve_left ha') ha'
Finset.prod_union_eq_right
theorem
[DecidableEq ι] (hs : ∀ a ∈ s₁, a ∉ s₂ → f a = 1) : ∏ a ∈ s₁ ∪ s₂, f a = ∏ a ∈ s₂, f a
@[to_additive]
lemma prod_union_eq_right [DecidableEq ι] (hs : ∀ a ∈ s₁, a ∉ s₂ → f a = 1) :
∏ a ∈ s₁ ∪ s₂, f a = ∏ a ∈ s₂, f a := by rw [union_comm, prod_union_eq_left hs]
Finset.prod_congr_of_eq_on_inter
theorem
{ι M : Type*} {s₁ s₂ : Finset ι} {f g : ι → M} [CommMonoid M] (h₁ : ∀ a ∈ s₁, a ∉ s₂ → f a = 1)
(h₂ : ∀ a ∈ s₂, a ∉ s₁ → g a = 1) (h : ∀ a ∈ s₁, a ∈ s₂ → f a = g a) : ∏ a ∈ s₁, f a = ∏ a ∈ s₂, g a
/-- The products of two functions `f g : ι → M` over finite sets `s₁ s₂ : Finset ι`
are equal if the functions agree on `s₁ ∩ s₂`, `f = 1` and `g = 1` on the respective
set differences. -/
@[to_additive "The sum of two functions `f g : ι → M` over finite sets `s₁ s₂ : Finset ι`
are equal if the functions agree on `s₁ ∩ s₂`, `f = 0` and `g = 0` on the respective
set differences."]
lemma prod_congr_of_eq_on_inter {ι M : Type*} {s₁ s₂ : Finset ι} {f g : ι → M} [CommMonoid M]
(h₁ : ∀ a ∈ s₁, a ∉ s₂ → f a = 1) (h₂ : ∀ a ∈ s₂, a ∉ s₁ → g a = 1)
(h : ∀ a ∈ s₁, a ∈ s₂ → f a = g a) :
∏ a ∈ s₁, f a = ∏ a ∈ s₂, g a := by
classical
conv_lhs => rw [← sdiff_union_inter s₁ s₂, prod_union_eq_right (by aesop)]
conv_rhs => rw [← sdiff_union_inter s₂ s₁, prod_union_eq_right (by aesop), inter_comm]
exact prod_congr rfl (by simpa)
Finset.prod_eq_mul_of_mem
theorem
{s : Finset ι} {f : ι → M} (a b : ι) (ha : a ∈ s) (hb : b ∈ s) (hn : a ≠ b) (h₀ : ∀ c ∈ s, c ≠ a ∧ c ≠ b → f c = 1) :
∏ x ∈ s, f x = f a * f b
@[to_additive]
theorem prod_eq_mul_of_mem {s : Finset ι} {f : ι → M} (a b : ι) (ha : a ∈ s) (hb : b ∈ s)
(hn : a ≠ b) (h₀ : ∀ c ∈ s, c ≠ a ∧ c ≠ b → f c = 1) : ∏ x ∈ s, f x = f a * f b := by
haveI := Classical.decEq ι; let s' := ({a, b} : Finset ι)
have hu : s' ⊆ s := by
refine insert_subset_iff.mpr ?_
apply And.intro ha
apply singleton_subset_iff.mpr hb
have hf : ∀ c ∈ s, c ∉ s' → f c = 1 := by
intro c hc hcs
apply h₀ c hc
apply not_or.mp
intro hab
apply hcs
rw [mem_insert, mem_singleton]
exact hab
rw [← prod_subset hu hf]
exact Finset.prod_pair hn
Finset.prod_eq_mul
theorem
{s : Finset ι} {f : ι → M} (a b : ι) (hn : a ≠ b) (h₀ : ∀ c ∈ s, c ≠ a ∧ c ≠ b → f c = 1) (ha : a ∉ s → f a = 1)
(hb : b ∉ s → f b = 1) : ∏ x ∈ s, f x = f a * f b
@[to_additive]
theorem prod_eq_mul {s : Finset ι} {f : ι → M} (a b : ι) (hn : a ≠ b)
(h₀ : ∀ c ∈ s, c ≠ a ∧ c ≠ b → f c = 1) (ha : a ∉ s → f a = 1) (hb : b ∉ s → f b = 1) :
∏ x ∈ s, f x = f a * f b := by
haveI := Classical.decEq ι; by_cases h₁ : a ∈ s <;> by_cases h₂ : b ∈ s
· exact prod_eq_mul_of_mem a b h₁ h₂ hn h₀
· rw [hb h₂, mul_one]
apply prod_eq_single_of_mem a h₁
exact fun c hc hca => h₀ c hc ⟨hca, ne_of_mem_of_not_mem hc h₂⟩
· rw [ha h₁, one_mul]
apply prod_eq_single_of_mem b h₂
exact fun c hc hcb => h₀ c hc ⟨ne_of_mem_of_not_mem hc h₁, hcb⟩
· rw [ha h₁, hb h₂, mul_one]
exact
_root_.trans
(prod_congr rfl fun c hc =>
h₀ c hc ⟨ne_of_mem_of_not_mem hc h₁, ne_of_mem_of_not_mem hc h₂⟩)
prod_const_one
Finset.prod_subtype_eq_prod_filter
theorem
(f : ι → M) {p : ι → Prop} [DecidablePred p] : ∏ x ∈ s.subtype p, f x = ∏ x ∈ s with p x, f x
/-- A product over `s.subtype p` equals one over `{x ∈ s | p x}`. -/
@[to_additive (attr := simp)
"A sum over `s.subtype p` equals one over `{x ∈ s | p x}`."]
theorem prod_subtype_eq_prod_filter (f : ι → M) {p : ι → Prop} [DecidablePred p] :
∏ x ∈ s.subtype p, f x = ∏ x ∈ s with p x, f x := by
have := prod_map (s.subtype p) (Function.Embedding.subtype _) f
simp_all
Finset.prod_subtype_of_mem
theorem
(f : ι → M) {p : ι → Prop} [DecidablePred p] (h : ∀ x ∈ s, p x) : ∏ x ∈ s.subtype p, f x = ∏ x ∈ s, f x
/-- If all elements of a `Finset` satisfy the predicate `p`, a product
over `s.subtype p` equals that product over `s`. -/
@[to_additive "If all elements of a `Finset` satisfy the predicate `p`, a sum
over `s.subtype p` equals that sum over `s`."]
theorem prod_subtype_of_mem (f : ι → M) {p : ι → Prop} [DecidablePred p] (h : ∀ x ∈ s, p x) :
∏ x ∈ s.subtype p, f x = ∏ x ∈ s, f x := by
rw [prod_subtype_eq_prod_filter, filter_true_of_mem]
simpa using h
Finset.prod_subtype_map_embedding
theorem
{p : ι → Prop} {s : Finset { x // p x }} {f : { x // p x } → M} {g : ι → M}
(h : ∀ x : { x // p x }, x ∈ s → g x = f x) : (∏ x ∈ s.map (Function.Embedding.subtype _), g x) = ∏ x ∈ s, f x
/-- A product of a function over a `Finset` in a subtype equals a
product in the main type of a function that agrees with the first
function on that `Finset`. -/
@[to_additive "A sum of a function over a `Finset` in a subtype equals a
sum in the main type of a function that agrees with the first
function on that `Finset`."]
theorem prod_subtype_map_embedding {p : ι → Prop} {s : Finset { x // p x }} {f : { x // p x } → M}
{g : ι → M} (h : ∀ x : { x // p x }, x ∈ s → g x = f x) :
(∏ x ∈ s.map (Function.Embedding.subtype _), g x) = ∏ x ∈ s, f x := by
rw [Finset.prod_map]
exact Finset.prod_congr rfl h
Finset.prod_coe_sort
theorem
: ∏ i : s, f i = ∏ i ∈ s, f i
@[to_additive]
theorem prod_coe_sort : ∏ i : s, f i = ∏ i ∈ s, f i := prod_attach _ _
Finset.prod_finset_coe
theorem
(f : ι → M) (s : Finset ι) : (∏ i : (s : Set ι), f i) = ∏ i ∈ s, f i
@[to_additive]
theorem prod_finset_coe (f : ι → M) (s : Finset ι) : (∏ i : (s : Set ι), f i) = ∏ i ∈ s, f i :=
prod_coe_sort s f
Finset.prod_subtype
theorem
{p : ι → Prop} {F : Fintype (Subtype p)} (s : Finset ι) (h : ∀ x, x ∈ s ↔ p x) (f : ι → M) :
∏ a ∈ s, f a = ∏ a : Subtype p, f a
@[to_additive]
theorem prod_subtype {p : ι → Prop} {F : Fintype (Subtype p)} (s : Finset ι) (h : ∀ x, x ∈ sx ∈ s ↔ p x)
(f : ι → M) : ∏ a ∈ s, f a = ∏ a : Subtype p, f a := by
have : (· ∈ s) = p := Set.ext h
subst p
rw [← prod_coe_sort]
congr!
Finset.prod_set_coe
theorem
(s : Set ι) [Fintype s] : (∏ i : s, f i) = ∏ i ∈ s.toFinset, f i
@[to_additive]
theorem prod_set_coe (s : Set ι) [Fintype s] : (∏ i : s, f i) = ∏ i ∈ s.toFinset, f i :=
(Finset.prod_subtype s.toFinset (fun _ ↦ Set.mem_toFinset) f).symm
Finset.prod_congr_set
theorem
[Fintype ι] (s : Set ι) [DecidablePred (· ∈ s)] (f : ι → M) (g : s → M) (w : ∀ x (hx : x ∈ s), f x = g ⟨x, hx⟩)
(w' : ∀ x ∉ s, f x = 1) : ∏ i, f i = ∏ i, g i
/-- The product of a function `g` defined only on a set `s` is equal to
the product of a function `f` defined everywhere,
as long as `f` and `g` agree on `s`, and `f = 1` off `s`. -/
@[to_additive "The sum of a function `g` defined only on a set `s` is equal to
the sum of a function `f` defined everywhere,
as long as `f` and `g` agree on `s`, and `f = 0` off `s`."]
theorem prod_congr_set [Fintype ι] (s : Set ι) [DecidablePred (· ∈ s)] (f : ι → M) (g : s → M)
(w : ∀ x (hx : x ∈ s), f x = g ⟨x, hx⟩) (w' : ∀ x ∉ s, f x = 1) : ∏ i, f i = ∏ i, g i := by
rw [← prod_subset s.toFinset.subset_univ (by simpa), prod_subtype (p := (· ∈ s)) _ (by simp)]
congr! with ⟨x, h⟩
exact w x h
Finset.prod_extend_by_one
theorem
[DecidableEq ι] (s : Finset ι) (f : ι → M) : ∏ i ∈ s, (if i ∈ s then f i else 1) = ∏ i ∈ s, f i
@[to_additive]
theorem prod_extend_by_one [DecidableEq ι] (s : Finset ι) (f : ι → M) :
∏ i ∈ s, (if i ∈ s then f i else 1) = ∏ i ∈ s, f i :=
(prod_congr rfl) fun _i hi => if_pos hi
Finset.prod_eq_prod_extend
theorem
(f : s → M) : ∏ x, f x = ∏ x ∈ s, Subtype.val.extend f 1 x
@[to_additive]
theorem prod_eq_prod_extend (f : s → M) : ∏ x, f x = ∏ x ∈ s, Subtype.val.extend f 1 x := by
rw [univ_eq_attach, ← Finset.prod_attach s]
congr with ⟨x, hx⟩
rw [Subtype.val_injective.extend_apply]
Finset.prod_bij_ne_one
theorem
{s : Finset ι} {t : Finset κ} {f : ι → M} {g : κ → M} (i : ∀ a ∈ s, f a ≠ 1 → κ) (hi : ∀ a h₁ h₂, i a h₁ h₂ ∈ t)
(i_inj : ∀ a₁ h₁₁ h₁₂ a₂ h₂₁ h₂₂, i a₁ h₁₁ h₁₂ = i a₂ h₂₁ h₂₂ → a₁ = a₂)
(i_surj : ∀ b ∈ t, g b ≠ 1 → ∃ a h₁ h₂, i a h₁ h₂ = b) (h : ∀ a h₁ h₂, f a = g (i a h₁ h₂)) :
∏ x ∈ s, f x = ∏ x ∈ t, g x
@[to_additive]
theorem prod_bij_ne_one {s : Finset ι} {t : Finset κ} {f : ι → M} {g : κ → M}
(i : ∀ a ∈ s, f a ≠ 1 → κ) (hi : ∀ a h₁ h₂, i a h₁ h₂ ∈ t)
(i_inj : ∀ a₁ h₁₁ h₁₂ a₂ h₂₁ h₂₂, i a₁ h₁₁ h₁₂ = i a₂ h₂₁ h₂₂ → a₁ = a₂)
(i_surj : ∀ b ∈ t, g b ≠ 1 → ∃ a h₁ h₂, i a h₁ h₂ = b) (h : ∀ a h₁ h₂, f a = g (i a h₁ h₂)) :
∏ x ∈ s, f x = ∏ x ∈ t, g x := by
classical
calc
∏ x ∈ s, f x = ∏ x ∈ s with f x ≠ 1, f x := by rw [prod_filter_ne_one]
_ = ∏ x ∈ t with g x ≠ 1, g x :=
prod_bij (fun a ha => i a (mem_filter.mp ha).1 <| by simpa using (mem_filter.mp ha).2)
?_ ?_ ?_ ?_
_ = ∏ x ∈ t, g x := prod_filter_ne_one _
· intros a ha
refine (mem_filter.mp ha).elim ?_
intros h₁ h₂
refine (mem_filter.mpr ⟨hi a h₁ _, ?_⟩)
specialize h a h₁ fun H ↦ by rw [H] at h₂; simp at h₂
rwa [← h]
· intros a₁ ha₁ a₂ ha₂
refine (mem_filter.mp ha₁).elim fun _ha₁₁ _ha₁₂ ↦ ?_
refine (mem_filter.mp ha₂).elim fun _ha₂₁ _ha₂₂ ↦ ?_
apply i_inj
· intros b hb
refine (mem_filter.mp hb).elim fun h₁ h₂ ↦ ?_
obtain ⟨a, ha₁, ha₂, eq⟩ := i_surj b h₁ fun H ↦ by rw [H] at h₂; simp at h₂
exact ⟨a, mem_filter.mpr ⟨ha₁, ha₂⟩, eq⟩
· refine (fun a ha => (mem_filter.mp ha).elim fun h₁ h₂ ↦ ?_)
exact h a h₁ fun H ↦ by rw [H] at h₂; simp at h₂
Finset.exists_ne_one_of_prod_ne_one
theorem
(h : ∏ x ∈ s, f x ≠ 1) : ∃ a ∈ s, f a ≠ 1
@[to_additive]
theorem exists_ne_one_of_prod_ne_one (h : ∏ x ∈ s, f x∏ x ∈ s, f x ≠ 1) : ∃ a ∈ s, f a ≠ 1 := by
classical
rw [← prod_filter_ne_one] at h
rcases nonempty_of_prod_ne_one h with ⟨x, hx⟩
exact ⟨x, (mem_filter.1 hx).1, by simpa using (mem_filter.1 hx).2⟩
Finset.prod_range_succ_comm
theorem
(f : ℕ → M) (n : ℕ) : (∏ x ∈ range (n + 1), f x) = f n * ∏ x ∈ range n, f x
@[to_additive]
theorem prod_range_succ_comm (f : ℕ → M) (n : ℕ) :
(∏ x ∈ range (n + 1), f x) = f n * ∏ x ∈ range n, f x := by
rw [range_succ, prod_insert not_mem_range_self]
Finset.prod_range_succ
theorem
(f : ℕ → M) (n : ℕ) : (∏ x ∈ range (n + 1), f x) = (∏ x ∈ range n, f x) * f n
@[to_additive]
theorem prod_range_succ (f : ℕ → M) (n : ℕ) :
(∏ x ∈ range (n + 1), f x) = (∏ x ∈ range n, f x) * f n := by
simp only [mul_comm, prod_range_succ_comm]
Finset.prod_range_succ'
theorem
(f : ℕ → M) : ∀ n : ℕ, (∏ k ∈ range (n + 1), f k) = (∏ k ∈ range n, f (k + 1)) * f 0
@[to_additive]
theorem prod_range_succ' (f : ℕ → M) :
∀ n : ℕ, (∏ k ∈ range (n + 1), f k) = (∏ k ∈ range n, f (k + 1)) * f 0
| 0 => prod_range_succ _ _
| n + 1 => by rw [prod_range_succ _ n, mul_right_comm, ← prod_range_succ' _ n, prod_range_succ]
Finset.eventually_constant_prod
theorem
{u : ℕ → M} {N : ℕ} (hu : ∀ n ≥ N, u n = 1) {n : ℕ} (hn : N ≤ n) : (∏ k ∈ range n, u k) = ∏ k ∈ range N, u k
@[to_additive]
theorem eventually_constant_prod {u : ℕ → M} {N : ℕ} (hu : ∀ n ≥ N, u n = 1) {n : ℕ} (hn : N ≤ n) :
(∏ k ∈ range n, u k) = ∏ k ∈ range N, u k := by
obtain ⟨m, rfl : n = N + m⟩ := Nat.exists_eq_add_of_le hn
clear hn
induction m with
| zero => simp
| succ m hm => simp [← add_assoc, prod_range_succ, hm, hu]
Finset.prod_range_add
theorem
(f : ℕ → M) (n m : ℕ) : (∏ x ∈ range (n + m), f x) = (∏ x ∈ range n, f x) * ∏ x ∈ range m, f (n + x)
@[to_additive]
theorem prod_range_add (f : ℕ → M) (n m : ℕ) :
(∏ x ∈ range (n + m), f x) = (∏ x ∈ range n, f x) * ∏ x ∈ range m, f (n + x) := by
induction m with
| zero => simp
| succ m hm => rw [Nat.add_succ, prod_range_succ, prod_range_succ, hm, mul_assoc]
Finset.prod_range_one
theorem
(f : ℕ → M) : ∏ k ∈ range 1, f k = f 0
@[to_additive sum_range_one]
theorem prod_range_one (f : ℕ → M) : ∏ k ∈ range 1, f k = f 0 := by
rw [range_one, prod_singleton]
Finset.prod_list_map_count
theorem
[DecidableEq ι] (l : List ι) (f : ι → M) : (l.map f).prod = ∏ m ∈ l.toFinset, f m ^ l.count m
@[to_additive]
theorem prod_list_map_count [DecidableEq ι] (l : List ι) (f : ι → M) :
(l.map f).prod = ∏ m ∈ l.toFinset, f m ^ l.count m := by
induction l with
| nil => simp only [map_nil, prod_nil, count_nil, pow_zero, prod_const_one]
| cons a s IH =>
simp only [List.map, List.prod_cons, toFinset_cons, IH]
by_cases has : a ∈ s.toFinset
· rw [insert_eq_of_mem has, ← insert_erase has, prod_insert (not_mem_erase _ _),
prod_insert (not_mem_erase _ _), ← mul_assoc, count_cons_self, pow_succ']
congr 1
refine prod_congr rfl fun x hx => ?_
rw [count_cons_of_ne (ne_of_mem_erase hx).symm]
rw [prod_insert has, count_cons_self, count_eq_zero_of_not_mem (mt mem_toFinset.2 has), pow_one]
congr 1
refine prod_congr rfl fun x hx => ?_
rw [count_cons_of_ne]
rintro rfl
exact has hx
Finset.prod_list_count
theorem
[DecidableEq M] (s : List M) : s.prod = ∏ m ∈ s.toFinset, m ^ s.count m
@[to_additive]
theorem prod_list_count [DecidableEq M] (s : List M) :
s.prod = ∏ m ∈ s.toFinset, m ^ s.count m := by simpa using prod_list_map_count s id
Finset.prod_list_count_of_subset
theorem
[DecidableEq M] (m : List M) (s : Finset M) (hs : m.toFinset ⊆ s) : m.prod = ∏ i ∈ s, i ^ m.count i
@[to_additive]
theorem prod_list_count_of_subset [DecidableEq M] (m : List M) (s : Finset M)
(hs : m.toFinset ⊆ s) : m.prod = ∏ i ∈ s, i ^ m.count i := by
rw [prod_list_count]
refine prod_subset hs fun x _ hx => ?_
rw [mem_toFinset] at hx
rw [count_eq_zero_of_not_mem hx, pow_zero]
Finset.prod_multiset_map_count
theorem
[DecidableEq ι] (s : Multiset ι) {M : Type*} [CommMonoid M] (f : ι → M) :
(s.map f).prod = ∏ m ∈ s.toFinset, f m ^ s.count m
@[to_additive]
theorem prod_multiset_map_count [DecidableEq ι] (s : Multiset ι) {M : Type*} [CommMonoid M]
(f : ι → M) : (s.map f).prod = ∏ m ∈ s.toFinset, f m ^ s.count m := by
refine Quot.induction_on s fun l => ?_
simp [prod_list_map_count l f]
Finset.prod_multiset_count
theorem
[DecidableEq M] (s : Multiset M) : s.prod = ∏ m ∈ s.toFinset, m ^ s.count m
@[to_additive]
theorem prod_multiset_count [DecidableEq M] (s : Multiset M) :
s.prod = ∏ m ∈ s.toFinset, m ^ s.count m := by
convert prod_multiset_map_count s id
rw [Multiset.map_id]
Finset.prod_multiset_count_of_subset
theorem
[DecidableEq M] (m : Multiset M) (s : Finset M) (hs : m.toFinset ⊆ s) : m.prod = ∏ i ∈ s, i ^ m.count i
@[to_additive]
theorem prod_multiset_count_of_subset [DecidableEq M] (m : Multiset M) (s : Finset M)
(hs : m.toFinset ⊆ s) : m.prod = ∏ i ∈ s, i ^ m.count i := by
revert hs
refine Quot.induction_on m fun l => ?_
simp only [quot_mk_to_coe'', prod_coe, coe_count]
apply prod_list_count_of_subset l s
Finset.prod_range_induction
theorem
(f s : ℕ → M) (base : s 0 = 1) (step : ∀ n, s (n + 1) = s n * f n) (n : ℕ) : ∏ k ∈ Finset.range n, f k = s n
/-- For any product along `{0, ..., n - 1}` of a commutative-monoid-valued function, we can verify
that it's equal to a different function just by checking ratios of adjacent terms.
This is a multiplicative discrete analogue of the fundamental theorem of calculus. -/
@[to_additive "For any sum along `{0, ..., n - 1}` of a commutative-monoid-valued function, we can
verify that it's equal to a different function just by checking differences of adjacent terms.
This is a discrete analogue of the fundamental theorem of calculus."]
theorem prod_range_induction (f s : ℕ → M) (base : s 0 = 1)
(step : ∀ n, s (n + 1) = s n * f n) (n : ℕ) :
∏ k ∈ Finset.range n, f k = s n := by
induction n with
| zero => rw [Finset.prod_range_zero, base]
| succ k hk => simp only [hk, Finset.prod_range_succ, step, mul_comm]
Finset.prod_const
theorem
(b : M) : ∏ _x ∈ s, b = b ^ #s
@[to_additive (attr := simp)]
theorem prod_const (b : M) : ∏ _x ∈ s, b = b ^ #s :=
(congr_arg _ <| s.val.map_const b).trans <| Multiset.prod_replicate #s b
Finset.prod_eq_pow_card
theorem
{b : M} (hf : ∀ a ∈ s, f a = b) : ∏ a ∈ s, f a = b ^ #s
@[to_additive sum_eq_card_nsmul]
theorem prod_eq_pow_card {b : M} (hf : ∀ a ∈ s, f a = b) : ∏ a ∈ s, f a = b ^ #s :=
(prod_congr rfl hf).trans <| prod_const _
Finset.pow_card_mul_prod
theorem
{b : M} : b ^ #s * ∏ a ∈ s, f a = ∏ a ∈ s, b * f a
@[to_additive card_nsmul_add_sum]
theorem pow_card_mul_prod {b : M} : b ^ #s * ∏ a ∈ s, f a = ∏ a ∈ s, b * f a :=
(Finset.prod_const b).symm ▸ prod_mul_distrib.symm
Finset.prod_mul_pow_card
theorem
{b : M} : (∏ a ∈ s, f a) * b ^ #s = ∏ a ∈ s, f a * b
@[to_additive sum_add_card_nsmul]
theorem prod_mul_pow_card {b : M} : (∏ a ∈ s, f a) * b ^ #s = ∏ a ∈ s, f a * b :=
(Finset.prod_const b).symm ▸ prod_mul_distrib.symm
Finset.pow_eq_prod_const
theorem
(b : M) : ∀ n, b ^ n = ∏ _k ∈ range n, b
@[to_additive]
theorem pow_eq_prod_const (b : M) : ∀ n, b ^ n = ∏ _k ∈ range n, b := by simp
Finset.prod_pow_eq_pow_sum
theorem
(s : Finset ι) (f : ι → ℕ) (a : M) : ∏ i ∈ s, a ^ f i = a ^ ∑ i ∈ s, f i
@[to_additive sum_nsmul_assoc]
lemma prod_pow_eq_pow_sum (s : Finset ι) (f : ι → ℕ) (a : M) :
∏ i ∈ s, a ^ f i = a ^ ∑ i ∈ s, f i :=
cons_induction (by simp) (fun _ _ _ _ ↦ by simp [prod_cons, sum_cons, pow_add, *]) s
Finset.prod_flip
theorem
{n : ℕ} (f : ℕ → M) : (∏ r ∈ range (n + 1), f (n - r)) = ∏ k ∈ range (n + 1), f k
@[to_additive]
theorem prod_flip {n : ℕ} (f : ℕ → M) :
(∏ r ∈ range (n + 1), f (n - r)) = ∏ k ∈ range (n + 1), f k := by
induction n with
| zero => rw [prod_range_one, prod_range_one]
| succ n ih =>
rw [prod_range_succ', prod_range_succ _ (Nat.succ n)]
simp [← ih]
Finset.prod_involution
theorem
(g : ∀ a ∈ s, ι) (hg₁ : ∀ a ha, f a * f (g a ha) = 1) (hg₃ : ∀ a ha, f a ≠ 1 → g a ha ≠ a) (g_mem : ∀ a ha, g a ha ∈ s)
(hg₄ : ∀ a ha, g (g a ha) (g_mem a ha) = a) : ∏ x ∈ s, f x = 1
/-- The difference with `Finset.prod_ninvolution` is that the involution is allowed to use
membership of the domain of the product, rather than being a non-dependent function. -/
@[to_additive "The difference with `Finset.sum_ninvolution` is that the involution is allowed to use
membership of the domain of the sum, rather than being a non-dependent function."]
lemma prod_involution (g : ∀ a ∈ s, ι) (hg₁ : ∀ a ha, f a * f (g a ha) = 1)
(hg₃ : ∀ a ha, f a ≠ 1 → g a ha ≠ a)
(g_mem : ∀ a ha, g a ha ∈ s) (hg₄ : ∀ a ha, g (g a ha) (g_mem a ha) = a) :
∏ x ∈ s, f x = 1 := by
classical
induction s using Finset.strongInduction with | H s ih => ?_
obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty
· simp
have : {x, g x hx}{x, g x hx} ⊆ s := by simp [insert_subset_iff, hx, g_mem]
suffices h : ∏ x ∈ s \ {x, g x hx}, f x = 1 by
rw [← prod_sdiff this, h, one_mul]
cases eq_or_ne (g x hx) x with
| inl hx' => simpa [hx'] using hg₃ x hx
| inr hx' => rw [prod_pair hx'.symm, hg₁]
suffices h₃ : ∀ a (ha : a ∈ s \ {x, g x hx}), g a (sdiff_subset ha) ∈ s \ {x, g x hx} from
ih (s \ {x, g x hx}) (ssubset_iff.2 ⟨x, by simp [insert_subset_iff, hx]⟩) _
(by simp [hg₁]) (fun _ _ => hg₃ _ _) h₃ (fun _ _ => hg₄ _ _)
simp only [mem_sdiff, mem_insert, mem_singleton, not_or, g_mem, true_and]
rintro a ⟨ha₁, ha₂, ha₃⟩
refine ⟨fun h => by simp [← h, hg₄] at ha₃, fun h => ?_⟩
have : g (g a ha₁) (g_mem _ _) = g (g x hx) (g_mem _ _) := by simp only [h]
exact ha₂ (by simpa [hg₄] using this)
Finset.prod_ninvolution
theorem
(g : ι → ι) (hg₁ : ∀ a, f a * f (g a) = 1) (hg₂ : ∀ a, f a ≠ 1 → g a ≠ a) (g_mem : ∀ a, g a ∈ s)
(hg₃ : ∀ a, g (g a) = a) : ∏ x ∈ s, f x = 1
/-- The difference with `Finset.prod_involution` is that the involution is a non-dependent function,
rather than being allowed to use membership of the domain of the product. -/
@[to_additive "The difference with `Finset.sum_involution` is that the involution is a non-dependent
function, rather than being allowed to use membership of the domain of the sum."]
lemma prod_ninvolution (g : ι → ι) (hg₁ : ∀ a, f a * f (g a) = 1) (hg₂ : ∀ a, f a ≠ 1 → g a ≠ a)
(g_mem : ∀ a, g a ∈ s) (hg₃ : ∀ a, g (g a) = a) : ∏ x ∈ s, f x = 1 :=
prod_involution (fun i _ => g i) (fun i _ => hg₁ i) (fun _ _ hi => hg₂ _ hi)
(fun i _ => g_mem i) (fun i _ => hg₃ i)
Finset.prod_comp
theorem
[DecidableEq κ] (f : κ → M) (g : ι → κ) : ∏ a ∈ s, f (g a) = ∏ b ∈ s.image g, f b ^ #({a ∈ s | g a = b})
/-- The product of the composition of functions `f` and `g`, is the product over `b ∈ s.image g` of
`f b` to the power of the cardinality of the fibre of `b`. See also `Finset.prod_image`. -/
@[to_additive "The sum of the composition of functions `f` and `g`, is the sum over `b ∈ s.image g`
of `f b` times of the cardinality of the fibre of `b`. See also `Finset.sum_image`."]
theorem prod_comp [DecidableEq κ] (f : κ → M) (g : ι → κ) :
∏ a ∈ s, f (g a) = ∏ b ∈ s.image g, f b ^ #{a ∈ s | g a = b} := by
simp_rw [← prod_const, prod_fiberwise_of_maps_to' fun _ ↦ mem_image_of_mem _]
Finset.prod_partition
theorem
(R : Setoid ι) [DecidableRel R.r] : ∏ x ∈ s, f x = ∏ xbar ∈ s.image (Quotient.mk _), ∏ y ∈ s with ⟦y⟧ = xbar, f y
/-- A product can be partitioned into a product of products, each equivalent under a setoid. -/
@[to_additive "A sum can be partitioned into a sum of sums, each equivalent under a setoid."]
theorem prod_partition (R : Setoid ι) [DecidableRel R.r] :
∏ x ∈ s, f x = ∏ xbar ∈ s.image (Quotient.mk _), ∏ y ∈ s with ⟦y⟧ = xbar, f y := by
refine (Finset.prod_image' f fun x _hx => ?_).symm
rfl
Finset.prod_cancels_of_partition_cancels
theorem
(R : Setoid ι) [DecidableRel R] (h : ∀ x ∈ s, ∏ a ∈ s with R a x, f a = 1) : ∏ x ∈ s, f x = 1
/-- If we can partition a product into subsets that cancel out, then the whole product cancels. -/
@[to_additive "If we can partition a sum into subsets that cancel out, then the whole sum cancels."]
theorem prod_cancels_of_partition_cancels (R : Setoid ι) [DecidableRel R]
(h : ∀ x ∈ s, ∏ a ∈ s with R a x, f a = 1) : ∏ x ∈ s, f x = 1 := by
rw [prod_partition R, ← Finset.prod_eq_one]
intro xbar xbar_in_s
obtain ⟨x, x_in_s, rfl⟩ := mem_image.mp xbar_in_s
simp only [← Quotient.eq] at h
exact h x x_in_s
Finset.mul_prod_erase
theorem
[DecidableEq ι] (s : Finset ι) (f : ι → M) {a : ι} (h : a ∈ s) : (f a * ∏ x ∈ s.erase a, f x) = ∏ x ∈ s, f x
/-- Taking a product over `s : Finset ι` is the same as multiplying the value on a single element
`f a` by the product of `s.erase a`.
See `Multiset.prod_map_erase` for the `Multiset` version. -/
@[to_additive "Taking a sum over `s : Finset ι` is the same as adding the value on a single element
`f a` to the sum over `s.erase a`.
See `Multiset.sum_map_erase` for the `Multiset` version."]
theorem mul_prod_erase [DecidableEq ι] (s : Finset ι) (f : ι → M) {a : ι} (h : a ∈ s) :
(f a * ∏ x ∈ s.erase a, f x) = ∏ x ∈ s, f x := by
rw [← prod_insert (not_mem_erase a s), insert_erase h]
Finset.prod_erase_mul
theorem
[DecidableEq ι] (s : Finset ι) (f : ι → M) {a : ι} (h : a ∈ s) : (∏ x ∈ s.erase a, f x) * f a = ∏ x ∈ s, f x
/-- A variant of `Finset.mul_prod_erase` with the multiplication swapped. -/
@[to_additive "A variant of `Finset.add_sum_erase` with the addition swapped."]
theorem prod_erase_mul [DecidableEq ι] (s : Finset ι) (f : ι → M) {a : ι} (h : a ∈ s) :
(∏ x ∈ s.erase a, f x) * f a = ∏ x ∈ s, f x := by rw [mul_comm, mul_prod_erase s f h]
Finset.prod_erase
theorem
[DecidableEq ι] (s : Finset ι) {f : ι → M} {a : ι} (h : f a = 1) : ∏ x ∈ s.erase a, f x = ∏ x ∈ s, f x
/-- If a function applied at a point is 1, a product is unchanged by
removing that point, if present, from a `Finset`. -/
@[to_additive "If a function applied at a point is 0, a sum is unchanged by
removing that point, if present, from a `Finset`."]
theorem prod_erase [DecidableEq ι] (s : Finset ι) {f : ι → M} {a : ι} (h : f a = 1) :
∏ x ∈ s.erase a, f x = ∏ x ∈ s, f x := by
rw [← sdiff_singleton_eq_erase]
refine prod_subset sdiff_subset fun x hx hnx => ?_
rw [sdiff_singleton_eq_erase] at hnx
rwa [eq_of_mem_of_not_mem_erase hx hnx]
Finset.prod_erase_lt_of_one_lt
theorem
{κ : Type*} [DecidableEq ι] [CommMonoid κ] [LT κ] [MulLeftStrictMono κ] {s : Finset ι} {d : ι} (hd : d ∈ s) {f : ι → κ}
(hdf : 1 < f d) : ∏ m ∈ s.erase d, f m < ∏ m ∈ s, f m
@[to_additive]
theorem prod_erase_lt_of_one_lt {κ : Type*} [DecidableEq ι] [CommMonoid κ] [LT κ]
[MulLeftStrictMono κ] {s : Finset ι} {d : ι} (hd : d ∈ s) {f : ι → κ}
(hdf : 1 < f d) : ∏ m ∈ s.erase d, f m < ∏ m ∈ s, f m := by
conv in ∏ m ∈ s, f m => rw [← Finset.insert_erase hd]
rw [Finset.prod_insert (Finset.not_mem_erase d s)]
exact lt_mul_of_one_lt_left' _ hdf
Finset.prod_mul_eq_prod_mul_of_exists
theorem
{s : Finset ι} {f : ι → M} {b₁ b₂ : M} (a : ι) (ha : a ∈ s) (h : f a * b₁ = f a * b₂) :
(∏ a ∈ s, f a) * b₁ = (∏ a ∈ s, f a) * b₂
@[to_additive]
lemma prod_mul_eq_prod_mul_of_exists {s : Finset ι} {f : ι → M} {b₁ b₂ : M}
(a : ι) (ha : a ∈ s) (h : f a * b₁ = f a * b₂) :
(∏ a ∈ s, f a) * b₁ = (∏ a ∈ s, f a) * b₂ := by
classical
rw [← insert_erase ha]
simp only [mem_erase, ne_eq, not_true_eq_false, false_and, not_false_eq_true, prod_insert]
rw [mul_assoc, mul_comm, mul_assoc, mul_comm b₁, h, ← mul_assoc, mul_comm _ (f a)]
Finset.prod_filter_of_pairwise_eq_one
theorem
[DecidableEq ι] {f : κ → ι} {g : ι → M} {n : κ} {I : Finset κ} (hn : n ∈ I)
(hf : (I : Set κ).Pairwise fun i j ↦ f i = f j → g (f i) = 1) : ∏ j ∈ I with f j = f n, g (f j) = g (f n)
@[to_additive]
lemma prod_filter_of_pairwise_eq_one [DecidableEq ι] {f : κ → ι} {g : ι → M} {n : κ} {I : Finset κ}
(hn : n ∈ I) (hf : (I : Set κ).Pairwise fun i j ↦ f i = f j → g (f i) = 1) :
∏ j ∈ I with f j = f n, g (f j) = g (f n) := by
classical
have h j (hj : j ∈ {i ∈ I | f i = f n}.erase n) : g (f j) = 1 := by
simp only [mem_erase, mem_filter] at hj
exact hf hj.2.1 hn hj.1 hj.2.2
rw [← mul_one (g (f n)), ← prod_eq_one h,
← mul_prod_erase {i ∈ I | f i = f n} (fun i ↦ g (f i)) <| mem_filter.mpr ⟨hn, by rfl⟩]
Finset.prod_image_of_pairwise_eq_one
theorem
[DecidableEq ι] {f : κ → ι} {g : ι → M} {I : Finset κ} (hf : (I : Set κ).Pairwise fun i j ↦ f i = f j → g (f i) = 1) :
∏ s ∈ I.image f, g s = ∏ i ∈ I, g (f i)
/-- A version of `Finset.prod_map` and `Finset.prod_image`, but we do not assume that `f` is
injective. Rather, we assume that the image of `f` on `I` only overlaps where `g (f i) = 1`.
The conclusion is the same as in `prod_image`. -/
@[to_additive (attr := simp)
"A version of `Finset.sum_map` and `Finset.sum_image`, but we do not assume that `f` is
injective. Rather, we assume that the image of `f` on `I` only overlaps where `g (f i) = 0`.
The conclusion is the same as in `sum_image`."]
lemma prod_image_of_pairwise_eq_one [DecidableEq ι] {f : κ → ι} {g : ι → M} {I : Finset κ}
(hf : (I : Set κ).Pairwise fun i j ↦ f i = f j → g (f i) = 1) :
∏ s ∈ I.image f, g s = ∏ i ∈ I, g (f i) := by
rw [prod_image']
exact fun n hnI => (prod_filter_of_pairwise_eq_one hnI hf).symm
Finset.prod_image_of_disjoint
theorem
[DecidableEq ι] [PartialOrder ι] [OrderBot ι] {f : κ → ι} {g : ι → M} (hg_bot : g ⊥ = 1) {I : Finset κ}
(hf_disj : (I : Set κ).PairwiseDisjoint f) : ∏ s ∈ I.image f, g s = ∏ i ∈ I, g (f i)
/-- A version of `Finset.prod_map` and `Finset.prod_image`, but we do not assume that `f` is
injective. Rather, we assume that the images of `f` are disjoint on `I`, and `g ⊥ = 1`. The
conclusion is the same as in `prod_image`. -/
@[to_additive (attr := simp)
"A version of `Finset.sum_map` and `Finset.sum_image`, but we do not assume that `f` is
injective. Rather, we assume that the images of `f` are disjoint on `I`, and `g ⊥ = 0`. The
conclusion is the same as in `sum_image`."]
lemma prod_image_of_disjoint [DecidableEq ι] [PartialOrder ι] [OrderBot ι] {f : κ → ι} {g : ι → M}
(hg_bot : g ⊥ = 1) {I : Finset κ} (hf_disj : (I : Set κ).PairwiseDisjoint f) :
∏ s ∈ I.image f, g s = ∏ i ∈ I, g (f i) := by
refine prod_image_of_pairwise_eq_one <| hf_disj.imp fun i j hdisj hfij ↦ ?_
rw [Function.onFun, ← hfij, disjoint_self] at hdisj
rw [hdisj, hg_bot]
Finset.prod_unique_nonempty
theorem
[Unique ι] (s : Finset ι) (f : ι → M) (h : s.Nonempty) : ∏ x ∈ s, f x = f default
@[to_additive]
theorem prod_unique_nonempty [Unique ι] (s : Finset ι) (f : ι → M) (h : s.Nonempty) :
∏ x ∈ s, f x = f default := by
rw [h.eq_singleton_default, Finset.prod_singleton]
Finset.prod_dvd_prod_of_dvd
theorem
(f g : ι → M) (h : ∀ i ∈ s, f i ∣ g i) : ∏ i ∈ s, f i ∣ ∏ i ∈ s, g i
lemma prod_dvd_prod_of_dvd (f g : ι → M) (h : ∀ i ∈ s, f i ∣ g i) :
∏ i ∈ s, f i∏ i ∈ s, f i ∣ ∏ i ∈ s, g i := by
induction s using Finset.cons_induction with
| empty => simp
| cons a T haT IH =>
rw [Finset.prod_cons, Finset.prod_cons]
rw [Finset.forall_mem_cons] at h
exact mul_dvd_mul h.1 <| IH h.2
Finset.prod_sdiff_eq_prod_sdiff_iff
theorem
: ∏ i ∈ s \ t, f i = ∏ i ∈ t \ s, f i ↔ ∏ i ∈ s, f i = ∏ i ∈ t, f i
@[to_additive]
lemma prod_sdiff_eq_prod_sdiff_iff :
∏ i ∈ s \ t, f i∏ i ∈ s \ t, f i = ∏ i ∈ t \ s, f i ↔ ∏ i ∈ s, f i = ∏ i ∈ t, f i :=
eq_comm.trans <| eq_iff_eq_of_mul_eq_mul <| by
rw [← prod_union disjoint_sdiff_self_left, ← prod_union disjoint_sdiff_self_left,
sdiff_union_self_eq_union, sdiff_union_self_eq_union, union_comm]
Finset.prod_sdiff_ne_prod_sdiff_iff
theorem
: ∏ i ∈ s \ t, f i ≠ ∏ i ∈ t \ s, f i ↔ ∏ i ∈ s, f i ≠ ∏ i ∈ t, f i
Finset.prod_insert_div
theorem
(ha : a ∉ s) (f : ι → G) : (∏ x ∈ insert a s, f x) / f a = ∏ x ∈ s, f x
@[to_additive]
lemma prod_insert_div (ha : a ∉ s) (f : ι → G) :
(∏ x ∈ insert a s, f x) / f a = ∏ x ∈ s, f x := by simp [ha]
Finset.prod_erase_eq_div
theorem
{a : ι} (h : a ∈ s) : ∏ x ∈ s.erase a, f x = (∏ x ∈ s, f x) / f a
@[to_additive (attr := simp)]
theorem prod_erase_eq_div {a : ι} (h : a ∈ s) : ∏ x ∈ s.erase a, f x = (∏ x ∈ s, f x) / f a := by
rw [eq_div_iff_mul_eq', prod_erase_mul _ _ h]
Finset.prod_range_div
theorem
(f : ℕ → G) (n : ℕ) : (∏ i ∈ range n, f (i + 1) / f i) = f n / f 0
/-- A telescoping product along `{0, ..., n - 1}` of a commutative group valued function reduces to
the ratio of the last and first factors. -/
@[to_additive "A telescoping sum along `{0, ..., n - 1}` of an additive commutative group valued
function reduces to the difference of the last and first terms."]
lemma prod_range_div (f : ℕ → G) (n : ℕ) : (∏ i ∈ range n, f (i + 1) / f i) = f n / f 0 := by
apply prod_range_induction <;> simp
Finset.prod_range_div'
theorem
(f : ℕ → G) (n : ℕ) : (∏ i ∈ range n, f i / f (i + 1)) = f 0 / f n
@[to_additive]
lemma prod_range_div' (f : ℕ → G) (n : ℕ) : (∏ i ∈ range n, f i / f (i + 1)) = f 0 / f n := by
apply prod_range_induction <;> simp
Finset.prod_range_add_div_prod_range
theorem
(f : ℕ → G) (n m : ℕ) : (∏ k ∈ range (n + m), f k) / ∏ k ∈ range n, f k = ∏ k ∈ Finset.range m, f (n + k)
@[to_additive]
lemma prod_range_add_div_prod_range (f : ℕ → G) (n m : ℕ) :
(∏ k ∈ range (n + m), f k) / ∏ k ∈ range n, f k = ∏ k ∈ Finset.range m, f (n + k) :=
div_eq_of_eq_mul' (prod_range_add f n m)
Finset.prod_sdiff_eq_div
theorem
(h : s₁ ⊆ s₂) : ∏ x ∈ s₂ \ s₁, f x = (∏ x ∈ s₂, f x) / ∏ x ∈ s₁, f x
@[to_additive (attr := simp)]
lemma prod_sdiff_eq_div (h : s₁ ⊆ s₂) : ∏ x ∈ s₂ \ s₁, f x = (∏ x ∈ s₂, f x) / ∏ x ∈ s₁, f x := by
rw [eq_div_iff_mul_eq', prod_sdiff h]
Finset.prod_sdiff_div_prod_sdiff
theorem
: (∏ x ∈ s₂ \ s₁, f x) / ∏ x ∈ s₁ \ s₂, f x = (∏ x ∈ s₂, f x) / ∏ x ∈ s₁, f x
@[to_additive]
theorem prod_sdiff_div_prod_sdiff :
(∏ x ∈ s₂ \ s₁, f x) / ∏ x ∈ s₁ \ s₂, f x = (∏ x ∈ s₂, f x) / ∏ x ∈ s₁, f x := by
simp [← Finset.prod_sdiff (@inf_le_left _ _ s₁ s₂), ← Finset.prod_sdiff (@inf_le_right _ _ s₁ s₂)]
Finset.sum_range_tsub
theorem
{f : ℕ → M} (h : Monotone f) (n : ℕ) : ∑ i ∈ range n, (f (i + 1) - f i) = f n - f 0
/-- A telescoping sum along `{0, ..., n-1}` of an `ℕ`-valued function reduces to the difference of
the last and first terms when the function we are summing is monotone. -/
lemma sum_range_tsub {f : ℕ → M} (h : Monotone f) (n : ℕ) :
∑ i ∈ range n, (f (i + 1) - f i) = f n - f 0 := by
apply sum_range_induction
case base => apply tsub_eq_of_eq_add; rw [zero_add]
case step =>
intro n
have h₁ : f n ≤ f (n + 1) := h (Nat.le_succ _)
have h₂ : f 0 ≤ f n := h (Nat.zero_le _)
rw [tsub_add_eq_add_tsub h₂, add_tsub_cancel_of_le h₁]
Finset.sum_tsub_distrib
theorem
(s : Finset ι) {f g : ι → M} (hfg : ∀ x ∈ s, g x ≤ f x) : ∑ x ∈ s, (f x - g x) = ∑ x ∈ s, f x - ∑ x ∈ s, g x
lemma sum_tsub_distrib (s : Finset ι) {f g : ι → M} (hfg : ∀ x ∈ s, g x ≤ f x) :
∑ x ∈ s, (f x - g x) = ∑ x ∈ s, f x - ∑ x ∈ s, g x := Multiset.sum_map_tsub _ hfg
Finset.card_eq_sum_ones
theorem
(s : Finset ι) : #s = ∑ _ ∈ s, 1
lemma card_eq_sum_ones (s : Finset ι) : #s = ∑ _ ∈ s, 1 := by simp
Finset.sum_const_nat
theorem
{m : ℕ} {f : ι → ℕ} (h₁ : ∀ x ∈ s, f x = m) : ∑ x ∈ s, f x = #s * m
theorem sum_const_nat {m : ℕ} {f : ι → ℕ} (h₁ : ∀ x ∈ s, f x = m) : ∑ x ∈ s, f x = #s * m := by
rw [← Nat.nsmul_eq_mul, ← sum_const]
apply sum_congr rfl h₁
Finset.sum_card_fiberwise_eq_card_filter
theorem
{κ : Type*} [DecidableEq κ] (s : Finset ι) (t : Finset κ) (g : ι → κ) :
∑ j ∈ t, #({i ∈ s | g i = j}) = #({i ∈ s | g i ∈ t})
lemma sum_card_fiberwise_eq_card_filter {κ : Type*} [DecidableEq κ] (s : Finset ι) (t : Finset κ)
(g : ι → κ) : ∑ j ∈ t, #{i ∈ s | g i = j} = #{i ∈ s | g i ∈ t} := by
simpa only [card_eq_sum_ones] using sum_fiberwise_eq_sum_filter _ _ _ _
Finset.card_disjiUnion
theorem
(s : Finset ι) (t : ι → Finset M) (h) : #(s.disjiUnion t h) = ∑ a ∈ s, #(t a)
@[simp]
theorem card_disjiUnion (s : Finset ι) (t : ι → Finset M) (h) :
#(s.disjiUnion t h) = ∑ a ∈ s, #(t a) :=
Multiset.card_bind _ _
Finset.card_biUnion
theorem
[DecidableEq M] {t : ι → Finset M} (h : s.toSet.PairwiseDisjoint t) : #(s.biUnion t) = ∑ u ∈ s, #(t u)
theorem card_biUnion [DecidableEq M] {t : ι → Finset M} (h : s.toSet.PairwiseDisjoint t) :
#(s.biUnion t) = ∑ u ∈ s, #(t u) := by simpa using sum_biUnion h (M := ℕ) (f := 1)
Finset.card_biUnion_le
theorem
[DecidableEq M] {s : Finset ι} {t : ι → Finset M} : #(s.biUnion t) ≤ ∑ a ∈ s, #(t a)
theorem card_biUnion_le [DecidableEq M] {s : Finset ι} {t : ι → Finset M} :
#(s.biUnion t) ≤ ∑ a ∈ s, #(t a) :=
haveI := Classical.decEq ι
Finset.induction_on s (by simp) fun a s has ih =>
calc
#((insert a s).biUnion t) ≤ #(t a) + #(s.biUnion t) := by
rw [biUnion_insert]; exact card_union_le ..
_ ≤ ∑ a ∈ insert a s, #(t a) := by rw [sum_insert has]; exact Nat.add_le_add_left ih _
Finset.card_eq_sum_card_fiberwise
theorem
[DecidableEq M] {f : ι → M} {s : Finset ι} {t : Finset M} (H : s.toSet.MapsTo f t) : #s = ∑ b ∈ t, #({a ∈ s | f a = b})
theorem card_eq_sum_card_fiberwise [DecidableEq M] {f : ι → M} {s : Finset ι} {t : Finset M}
(H : s.toSet.MapsTo f t) : #s = ∑ b ∈ t, #{a ∈ s | f a = b} := by
simp only [card_eq_sum_ones, sum_fiberwise_of_maps_to H]
Finset.card_eq_sum_card_image
theorem
[DecidableEq M] (f : ι → M) (s : Finset ι) : #s = ∑ b ∈ s.image f, #({a ∈ s | f a = b})
theorem card_eq_sum_card_image [DecidableEq M] (f : ι → M) (s : Finset ι) :
#s = ∑ b ∈ s.image f, #{a ∈ s | f a = b} :=
card_eq_sum_card_fiberwise fun _ => mem_image_of_mem _
Finset.mem_sum
theorem
{f : ι → Multiset M} (s : Finset ι) (b : M) : (b ∈ ∑ x ∈ s, f x) ↔ ∃ a ∈ s, b ∈ f a
theorem mem_sum {f : ι → Multiset M} (s : Finset ι) (b : M) :
(b ∈ ∑ x ∈ s, f x) ↔ ∃ a ∈ s, b ∈ f a := by
induction s using Finset.cons_induction with
| empty => simp
| cons a t hi ih => simp [sum_cons, ih, or_and_right, exists_or]
Fintype.prod_of_injective
theorem
(e : ι → κ) (he : Injective e) (f : ι → M) (g : κ → M) (h' : ∀ i ∉ Set.range e, g i = 1) (h : ∀ i, f i = g (e i)) :
∏ i, f i = ∏ j, g j
@[to_additive]
lemma prod_of_injective (e : ι → κ) (he : Injective e) (f : ι → M) (g : κ → M)
(h' : ∀ i ∉ Set.range e, g i = 1) (h : ∀ i, f i = g (e i)) : ∏ i, f i = ∏ j, g j :=
prod_of_injOn e he.injOn (by simp) (by simpa using h') (fun i _ ↦ h i)
Fintype.prod_fiberwise
theorem
[DecidableEq κ] (g : ι → κ) (f : ι → M) : ∏ j, ∏ i : { i // g i = j }, f i = ∏ i, f i
@[to_additive]
lemma prod_fiberwise [DecidableEq κ] (g : ι → κ) (f : ι → M) :
∏ j, ∏ i : {i // g i = j}, f i = ∏ i, f i := by
rw [← Finset.prod_fiberwise _ g f]
congr with j
exact (prod_subtype _ (by simp) _).symm
Fintype.prod_fiberwise'
theorem
[DecidableEq κ] (g : ι → κ) (f : κ → M) : ∏ j, ∏ _i : { i // g i = j }, f j = ∏ i, f (g i)
@[to_additive]
lemma prod_fiberwise' [DecidableEq κ] (g : ι → κ) (f : κ → M) :
∏ j, ∏ _i : {i // g i = j}, f j = ∏ i, f (g i) := by
rw [← Finset.prod_fiberwise' _ g f]
congr with j
exact (prod_subtype _ (by simp) fun _ ↦ _).symm
Fintype.prod_unique
theorem
[Unique ι] (f : ι → M) : ∏ x : ι, f x = f default
@[to_additive]
theorem prod_unique [Unique ι] (f : ι → M) : ∏ x : ι, f x = f default := by
rw [univ_unique, prod_singleton]
Fintype.prod_subsingleton
theorem
[Subsingleton ι] (f : ι → M) (a : ι) : ∏ x : ι, f x = f a
@[to_additive]
theorem prod_subsingleton [Subsingleton ι] (f : ι → M) (a : ι) : ∏ x : ι, f x = f a := by
have : Unique ι := uniqueOfSubsingleton a
rw [prod_unique f, Subsingleton.elim default a]
Fintype.prod_Prop
theorem
(f : Prop → M) : ∏ p, f p = f True * f False
@[to_additive] theorem prod_Prop (f : Prop → M) : ∏ p, f p = f True * f False := by simp
Fintype.prod_subtype_mul_prod_subtype
theorem
(p : ι → Prop) (f : ι → M) [DecidablePred p] : (∏ i : { x // p x }, f i) * ∏ i : { x // ¬p x }, f i = ∏ i, f i
@[to_additive]
theorem prod_subtype_mul_prod_subtype (p : ι → Prop) (f : ι → M) [DecidablePred p] :
(∏ i : { x // p x }, f i) * ∏ i : { x // ¬p x }, f i = ∏ i, f i := by
classical
let s := { x | p x }.toFinset
rw [← Finset.prod_subtype s, ← Finset.prod_subtype sᶜ]
· exact Finset.prod_mul_prod_compl _ _
· simp [s]
· simp [s]
Fintype.prod_subset
theorem
{s : Finset ι} {f : ι → M} (h : ∀ i, f i ≠ 1 → i ∈ s) : ∏ i ∈ s, f i = ∏ i, f i
@[to_additive] lemma prod_subset {s : Finset ι} {f : ι → M} (h : ∀ i, f i ≠ 1 → i ∈ s) :
∏ i ∈ s, f i = ∏ i, f i :=
Finset.prod_subset s.subset_univ <| by simpa [not_imp_comm (a := _ ∈ s)]
List.prod_toFinset
theorem
{M : Type*} [DecidableEq ι] [CommMonoid M] (f : ι → M) :
∀ {l : List ι} (_hl : l.Nodup), l.toFinset.prod f = (l.map f).prod
@[to_additive]
theorem prod_toFinset {M : Type*} [DecidableEq ι] [CommMonoid M] (f : ι → M) :
∀ {l : List ι} (_hl : l.Nodup), l.toFinset.prod f = (l.map f).prod
| [], _ => by simp
| a :: l, hl => by
let ⟨not_mem, hl⟩ := List.nodup_cons.mp hl
simp [Finset.prod_insert (mt List.mem_toFinset.mp not_mem), prod_toFinset _ hl]
List.sum_toFinset_count_eq_length
theorem
[DecidableEq ι] (l : List ι) : ∑ a ∈ l.toFinset, l.count a = l.length
@[simp]
theorem sum_toFinset_count_eq_length [DecidableEq ι] (l : List ι) :
∑ a ∈ l.toFinset, l.count a = l.length := by
simpa [List.map_const'] using (Finset.sum_list_map_count l fun _ => (1 : ℕ)).symm
Multiset.mem_sum
theorem
{s : Finset ι} {m : ι → Multiset ι} : a ∈ ∑ i ∈ s, m i ↔ ∃ i ∈ s, a ∈ m i
@[simp]
lemma mem_sum {s : Finset ι} {m : ι → Multiset ι} : a ∈ ∑ i ∈ s, m ia ∈ ∑ i ∈ s, m i ↔ ∃ i ∈ s, a ∈ m i := by
induction s using Finset.cons_induction <;> simp [*]
Multiset.toFinset_sum_count_eq
theorem
(s : Multiset ι) : ∑ a ∈ s.toFinset, s.count a = card s
theorem toFinset_sum_count_eq (s : Multiset ι) : ∑ a ∈ s.toFinset, s.count a = card s := by
simpa using (Finset.sum_multiset_map_count s (fun _ => (1 : ℕ))).symm
Multiset.sum_count_eq_card
theorem
{s : Finset ι} {m : Multiset ι} (hms : ∀ a ∈ m, a ∈ s) : ∑ a ∈ s, m.count a = card m
@[simp] lemma sum_count_eq_card {s : Finset ι} {m : Multiset ι} (hms : ∀ a ∈ m, a ∈ s) :
∑ a ∈ s, m.count a = card m := by
rw [← toFinset_sum_count_eq, ← Finset.sum_filter_ne_zero]
congr with a
simpa using hms a
Multiset.toFinset_sum_count_nsmul_eq
theorem
(s : Multiset ι) : ∑ a ∈ s.toFinset, s.count a • { a } = s
@[simp]
theorem toFinset_sum_count_nsmul_eq (s : Multiset ι) :
∑ a ∈ s.toFinset, s.count a • {a} = s := by
rw [← Finset.sum_multiset_map_count, Multiset.sum_map_singleton]
Multiset.exists_smul_of_dvd_count
theorem
(s : Multiset ι) {k : ℕ} (h : ∀ a : ι, a ∈ s → k ∣ Multiset.count a s) : ∃ u : Multiset ι, s = k • u
theorem exists_smul_of_dvd_count (s : Multiset ι) {k : ℕ}
(h : ∀ a : ι, a ∈ s → k ∣ Multiset.count a s) : ∃ u : Multiset ι, s = k • u := by
use ∑ a ∈ s.toFinset, (s.count a / k) • {a}
have h₂ :
(∑ x ∈ s.toFinset, k • (count x s / k) • ({x} : Multiset ι)) =
∑ x ∈ s.toFinset, count x s • {x} := by
apply Finset.sum_congr rfl
intro x hx
rw [← mul_nsmul', Nat.mul_div_cancel' (h x (mem_toFinset.mp hx))]
rw [← Finset.sum_nsmul, h₂, toFinset_sum_count_nsmul_eq]
Multiset.prod_sum
theorem
{ι : Type*} [CommMonoid M] (f : ι → Multiset M) (s : Finset ι) : (∑ x ∈ s, f x).prod = ∏ x ∈ s, (f x).prod
@[to_additive]
theorem prod_sum {ι : Type*} [CommMonoid M] (f : ι → Multiset M) (s : Finset ι) :
(∑ x ∈ s, f x).prod = ∏ x ∈ s, (f x).prod := by
induction s using Finset.cons_induction with
| empty => rw [Finset.sum_empty, Finset.prod_empty, Multiset.prod_zero]
| cons a s has ih => rw [Finset.sum_cons, Finset.prod_cons, Multiset.prod_add, ih]
IsUnit.prod_iff
theorem
[CommMonoid M] {f : ι → M} : IsUnit (∏ a ∈ s, f a) ↔ ∀ a ∈ s, IsUnit (f a)
@[to_additive (attr := simp)]
lemma IsUnit.prod_iff [CommMonoid M] {f : ι → M} :
IsUnitIsUnit (∏ a ∈ s, f a) ↔ ∀ a ∈ s, IsUnit (f a) := by
induction s using Finset.cons_induction with
| empty => simp
| cons a s ha hs => rw [Finset.prod_cons, IsUnit.mul_iff, hs, Finset.forall_mem_cons]
IsUnit.prod_univ_iff
theorem
[Fintype ι] [CommMonoid M] {f : ι → M} : IsUnit (∏ a, f a) ↔ ∀ a, IsUnit (f a)
@[to_additive]
lemma IsUnit.prod_univ_iff [Fintype ι] [CommMonoid M] {f : ι → M} :
IsUnitIsUnit (∏ a, f a) ↔ ∀ a, IsUnit (f a) := by simp
nat_abs_sum_le
theorem
(s : Finset ι) (f : ι → ℤ) : (∑ i ∈ s, f i).natAbs ≤ ∑ i ∈ s, (f i).natAbs
theorem nat_abs_sum_le (s : Finset ι) (f : ι → ℤ) :
(∑ i ∈ s, f i).natAbs ≤ ∑ i ∈ s, (f i).natAbs := by
induction s using Finset.cons_induction with
| empty => simp only [Finset.sum_empty, Int.natAbs_zero, le_refl]
| cons i s his IH =>
simp only [Finset.sum_cons, not_false_iff]
exact (Int.natAbs_add_le _ _).trans (Nat.add_le_add_left IH _)