Documentation

Mathlib.Algebra.BigOperators.Group.Finset.Defs

Big operators #

In this file we define products and sums indexed by finite sets (specifically, Finset).

Notation #

We introduce the following notation.

Let s be a Finset α, and f : α → β a function.

Implementation Notes #

The first arguments in all definitions and lemmas is the codomain of the function of the big operator. This is necessary for the heuristic in @[to_additive]. See the documentation of to_additive.attr for more information.

def Finset.prod {α : Type u_3} {β : Type u_4} [CommMonoid β] (s : Finset α) (f : αβ) :
β

∏ x ∈ s, f x is the product of f x as x ranges over the elements of the finite set s.

When the index type is a Fintype, the notation ∏ x, f x, is a shorthand for ∏ x ∈ Finset.univ, f x.

Equations
Instances For
    def Finset.sum {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (s : Finset α) (f : αβ) :
    β

    ∑ x ∈ s, f x is the sum of f x as x ranges over the elements of the finite set s.

    When the index type is a Fintype, the notation ∑ x, f x, is a shorthand for ∑ x ∈ Finset.univ, f x.

    Equations
    Instances For
      @[simp]
      theorem Finset.prod_mk {α : Type u_3} {β : Type u_4} [CommMonoid β] (s : Multiset α) (hs : s.Nodup) (f : αβ) :
      { val := s, nodup := hs }.prod f = (Multiset.map f s).prod
      @[simp]
      theorem Finset.sum_mk {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (s : Multiset α) (hs : s.Nodup) (f : αβ) :
      { val := s, nodup := hs }.sum f = (Multiset.map f s).sum
      @[simp]
      theorem Finset.prod_val {α : Type u_3} [CommMonoid α] (s : Finset α) :
      s.val.prod = s.prod id
      @[simp]
      theorem Finset.sum_val {α : Type u_3} [AddCommMonoid α] (s : Finset α) :
      s.val.sum = s.sum id

      A bigOpBinder is like an extBinder and has the form x, x : ty, or x pred where pred is a binderPred like < 2. Unlike extBinder, x is a term.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        A BigOperator binder in parentheses

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          A list of parenthesized binders

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            A single (unparenthesized) binder, or a list of parenthesized binders

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              def BigOperators.processBigOpBinder (processed : Array (Lean.Term × Lean.Term)) (binder : Lean.TSyntax `BigOperators.bigOpBinder) :

              Collects additional binder/Finset pairs for the given bigOpBinder. Note: this is not extensible at the moment, unlike the usual bigOpBinder expansions.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                def BigOperators.processBigOpBinders (binders : Lean.TSyntax `BigOperators.bigOpBinders) :

                Collects the binder/Finset pairs for the given bigOpBinders.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For

                  Collect the binderIdents into a ⟨...⟩ expression.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For

                    Collect the terms into a product of sets.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      • ∑ x, f x is notation for Finset.sum Finset.univ f. It is the sum of f x, where x ranges over the finite domain of f.
                      • ∑ x ∈ s, f x is notation for Finset.sum s f. It is the sum of f x, where x ranges over the finite set s (either a Finset or a Set with a Fintype instance).
                      • ∑ x ∈ s with p x, f x is notation for Finset.sum (Finset.filter p s) f.
                      • ∑ (x ∈ s) (y ∈ t), f x y is notation for Finset.sum (s ×ˢ t) (fun ⟨x, y⟩ ↦ f x y).

                      These support destructuring, for example ∑ ⟨x, y⟩ ∈ s ×ˢ t, f x y.

                      Notation: "∑" bigOpBinders* ("with" term)? "," term

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        • ∏ x, f x is notation for Finset.prod Finset.univ f. It is the product of f x, where x ranges over the finite domain of f.
                        • ∏ x ∈ s, f x is notation for Finset.prod s f. It is the product of f x, where x ranges over the finite set s (either a Finset or a Set with a Fintype instance).
                        • ∏ x ∈ s with p x, f x is notation for Finset.prod (Finset.filter p s) f.
                        • ∏ (x ∈ s) (y ∈ t), f x y is notation for Finset.prod (s ×ˢ t) (fun ⟨x, y⟩ ↦ f x y).

                        These support destructuring, for example ∏ ⟨x, y⟩ ∈ s ×ˢ t, f x y.

                        Notation: "∏" bigOpBinders* ("with" term)? "," term

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For

                          (Deprecated, use ∑ x ∈ s, f x) ∑ x in s, f x is notation for Finset.sum s f. It is the sum of f x, where x ranges over the finite set s.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For

                            (Deprecated, use ∏ x ∈ s, f x) ∏ x in s, f x is notation for Finset.prod s f. It is the product of f x, where x ranges over the finite set s.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For

                              Delaborator for Finset.prod. The pp.piBinderTypes option controls whether to show the domain type when the product is over Finset.univ.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For

                                Delaborator for Finset.sum. The pp.piBinderTypes option controls whether to show the domain type when the sum is over Finset.univ.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  theorem Finset.prod_eq_multiset_prod {α : Type u_3} {β : Type u_4} [CommMonoid β] (s : Finset α) (f : αβ) :
                                  xs, f x = (Multiset.map f s.val).prod
                                  theorem Finset.sum_eq_multiset_sum {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (s : Finset α) (f : αβ) :
                                  xs, f x = (Multiset.map f s.val).sum
                                  @[simp]
                                  theorem Finset.prod_map_val {α : Type u_3} {β : Type u_4} [CommMonoid β] (s : Finset α) (f : αβ) :
                                  (Multiset.map f s.val).prod = as, f a
                                  @[simp]
                                  theorem Finset.sum_map_val {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (s : Finset α) (f : αβ) :
                                  (Multiset.map f s.val).sum = as, f a
                                  @[simp]
                                  theorem Finset.sum_multiset_singleton {α : Type u_3} (s : Finset α) :
                                  xs, {x} = s.val
                                  @[simp]
                                  theorem map_prod {α : Type u_3} {β : Type u_4} {γ : Type u_5} [CommMonoid β] [CommMonoid γ] {G : Type u_6} [FunLike G β γ] [MonoidHomClass G β γ] (g : G) (f : αβ) (s : Finset α) :
                                  g (∏ xs, f x) = xs, g (f x)
                                  @[simp]
                                  theorem map_sum {α : Type u_3} {β : Type u_4} {γ : Type u_5} [AddCommMonoid β] [AddCommMonoid γ] {G : Type u_6} [FunLike G β γ] [AddMonoidHomClass G β γ] (g : G) (f : αβ) (s : Finset α) :
                                  g (∑ xs, f x) = xs, g (f x)
                                  @[simp]
                                  theorem Finset.prod_empty {α : Type u_3} {β : Type u_4} {f : αβ} [CommMonoid β] :
                                  x, f x = 1
                                  @[simp]
                                  theorem Finset.sum_empty {α : Type u_3} {β : Type u_4} {f : αβ} [AddCommMonoid β] :
                                  x, f x = 0
                                  theorem Finset.prod_of_isEmpty {α : Type u_3} {β : Type u_4} {f : αβ} [CommMonoid β] [IsEmpty α] (s : Finset α) :
                                  is, f i = 1
                                  theorem Finset.sum_of_isEmpty {α : Type u_3} {β : Type u_4} {f : αβ} [AddCommMonoid β] [IsEmpty α] (s : Finset α) :
                                  is, f i = 0
                                  @[deprecated Finset.prod_of_isEmpty (since := "2024-06-11")]
                                  theorem Finset.prod_of_empty {α : Type u_3} {β : Type u_4} {f : αβ} [CommMonoid β] [IsEmpty α] (s : Finset α) :
                                  is, f i = 1

                                  Alias of Finset.prod_of_isEmpty.

                                  @[deprecated Finset.sum_of_isEmpty (since := "2024-06-11")]
                                  theorem Finset.sum_of_empty {α : Type u_3} {β : Type u_4} {f : αβ} [AddCommMonoid β] [IsEmpty α] (s : Finset α) :
                                  is, f i = 0

                                  Alias of Finset.sum_of_isEmpty.

                                  @[simp]
                                  theorem Finset.prod_const_one {α : Type u_3} {β : Type u_4} {s : Finset α} [CommMonoid β] :
                                  _xs, 1 = 1
                                  @[simp]
                                  theorem Finset.sum_const_zero {α : Type u_3} {β : Type u_4} {s : Finset α} [AddCommMonoid β] :
                                  _xs, 0 = 0
                                  @[simp]
                                  theorem Finset.prod_map {α : Type u_3} {β : Type u_4} {γ : Type u_5} [CommMonoid β] (s : Finset α) (e : α γ) (f : γβ) :
                                  xFinset.map e s, f x = xs, f (e x)
                                  @[simp]
                                  theorem Finset.sum_map {α : Type u_3} {β : Type u_4} {γ : Type u_5} [AddCommMonoid β] (s : Finset α) (e : α γ) (f : γβ) :
                                  xFinset.map e s, f x = xs, f (e x)
                                  @[simp]
                                  theorem Finset.prod_to_list {α : Type u_3} {β : Type u_4} [CommMonoid β] (s : Finset α) (f : αβ) :
                                  (List.map f s.toList).prod = s.prod f
                                  @[simp]
                                  theorem Finset.sum_to_list {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (s : Finset α) (f : αβ) :
                                  (List.map f s.toList).sum = s.sum f
                                  theorem Equiv.Perm.prod_comp {α : Type u_3} {β : Type u_4} [CommMonoid β] (σ : Equiv.Perm α) (s : Finset α) (f : αβ) (hs : {a : α | σ a a} s) :
                                  xs, f (σ x) = xs, f x
                                  theorem Equiv.Perm.sum_comp {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (σ : Equiv.Perm α) (s : Finset α) (f : αβ) (hs : {a : α | σ a a} s) :
                                  xs, f (σ x) = xs, f x
                                  theorem Equiv.Perm.prod_comp' {α : Type u_3} {β : Type u_4} [CommMonoid β] (σ : Equiv.Perm α) (s : Finset α) (f : ααβ) (hs : {a : α | σ a a} s) :
                                  xs, f (σ x) x = xs, f x ((Equiv.symm σ) x)
                                  theorem Equiv.Perm.sum_comp' {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (σ : Equiv.Perm α) (s : Finset α) (f : ααβ) (hs : {a : α | σ a a} s) :
                                  xs, f (σ x) x = xs, f x ((Equiv.symm σ) x)
                                  theorem Finset.prod_bij {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [CommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : (a : ι) → a sκ) (hi : ∀ (a : ι) (ha : a s), i a ha t) (i_inj : ∀ (a₁ : ι) (ha₁ : a₁ s) (a₂ : ι) (ha₂ : a₂ s), i a₁ ha₁ = i a₂ ha₂a₁ = a₂) (i_surj : bt, ∃ (a : ι) (ha : a s), i a ha = b) (h : ∀ (a : ι) (ha : a s), f a = g (i a ha)) :
                                  xs, f x = xt, g x

                                  Reorder a product.

                                  The difference with Finset.prod_bij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

                                  The difference with Finset.prod_nbij is that the bijection is allowed to use membership of the domain of the product, rather than being a non-dependent function.

                                  theorem Finset.sum_bij {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [AddCommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : (a : ι) → a sκ) (hi : ∀ (a : ι) (ha : a s), i a ha t) (i_inj : ∀ (a₁ : ι) (ha₁ : a₁ s) (a₂ : ι) (ha₂ : a₂ s), i a₁ ha₁ = i a₂ ha₂a₁ = a₂) (i_surj : bt, ∃ (a : ι) (ha : a s), i a ha = b) (h : ∀ (a : ι) (ha : a s), f a = g (i a ha)) :
                                  xs, f x = xt, g x

                                  Reorder a sum.

                                  The difference with Finset.sum_bij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

                                  The difference with Finset.sum_nbij is that the bijection is allowed to use membership of the domain of the sum, rather than being a non-dependent function.

                                  theorem Finset.prod_bij' {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [CommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : (a : ι) → a sκ) (j : (a : κ) → a tι) (hi : ∀ (a : ι) (ha : a s), i a ha t) (hj : ∀ (a : κ) (ha : a t), j a ha s) (left_inv : ∀ (a : ι) (ha : a s), j (i a ha) = a) (right_inv : ∀ (a : κ) (ha : a t), i (j a ha) = a) (h : ∀ (a : ι) (ha : a s), f a = g (i a ha)) :
                                  xs, f x = xt, g x

                                  Reorder a product.

                                  The difference with Finset.prod_bij is that the bijection is specified with an inverse, rather than as a surjective injection.

                                  The difference with Finset.prod_nbij' is that the bijection and its inverse are allowed to use membership of the domains of the products, rather than being non-dependent functions.

                                  theorem Finset.sum_bij' {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [AddCommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : (a : ι) → a sκ) (j : (a : κ) → a tι) (hi : ∀ (a : ι) (ha : a s), i a ha t) (hj : ∀ (a : κ) (ha : a t), j a ha s) (left_inv : ∀ (a : ι) (ha : a s), j (i a ha) = a) (right_inv : ∀ (a : κ) (ha : a t), i (j a ha) = a) (h : ∀ (a : ι) (ha : a s), f a = g (i a ha)) :
                                  xs, f x = xt, g x

                                  Reorder a sum.

                                  The difference with Finset.sum_bij is that the bijection is specified with an inverse, rather than as a surjective injection.

                                  The difference with Finset.sum_nbij' is that the bijection and its inverse are allowed to use membership of the domains of the sums, rather than being non-dependent functions.

                                  theorem Finset.prod_nbij {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [CommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : ικ) (hi : as, i a t) (i_inj : Set.InjOn i s) (i_surj : Set.SurjOn i s t) (h : as, f a = g (i a)) :
                                  xs, f x = xt, g x

                                  Reorder a product.

                                  The difference with Finset.prod_nbij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

                                  The difference with Finset.prod_bij is that the bijection is a non-dependent function, rather than being allowed to use membership of the domain of the product.

                                  theorem Finset.sum_nbij {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [AddCommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : ικ) (hi : as, i a t) (i_inj : Set.InjOn i s) (i_surj : Set.SurjOn i s t) (h : as, f a = g (i a)) :
                                  xs, f x = xt, g x

                                  Reorder a sum.

                                  The difference with Finset.sum_nbij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

                                  The difference with Finset.sum_bij is that the bijection is a non-dependent function, rather than being allowed to use membership of the domain of the sum.

                                  theorem Finset.prod_nbij' {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [CommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : ικ) (j : κι) (hi : as, i a t) (hj : at, j a s) (left_inv : as, j (i a) = a) (right_inv : at, i (j a) = a) (h : as, f a = g (i a)) :
                                  xs, f x = xt, g x

                                  Reorder a product.

                                  The difference with Finset.prod_nbij is that the bijection is specified with an inverse, rather than as a surjective injection.

                                  The difference with Finset.prod_bij' is that the bijection and its inverse are non-dependent functions, rather than being allowed to use membership of the domains of the products.

                                  The difference with Finset.prod_equiv is that bijectivity is only required to hold on the domains of the products, rather than on the entire types.

                                  theorem Finset.sum_nbij' {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [AddCommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (i : ικ) (j : κι) (hi : as, i a t) (hj : at, j a s) (left_inv : as, j (i a) = a) (right_inv : at, i (j a) = a) (h : as, f a = g (i a)) :
                                  xs, f x = xt, g x

                                  Reorder a sum.

                                  The difference with Finset.sum_nbij is that the bijection is specified with an inverse, rather than as a surjective injection.

                                  The difference with Finset.sum_bij' is that the bijection and its inverse are non-dependent functions, rather than being allowed to use membership of the domains of the sums.

                                  The difference with Finset.sum_equiv is that bijectivity is only required to hold on the domains of the sums, rather than on the entire types.

                                  theorem Finset.prod_equiv {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [CommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (e : ι κ) (hst : ∀ (i : ι), i s e i t) (hfg : is, f i = g (e i)) :
                                  is, f i = it, g i

                                  Specialization of Finset.prod_nbij' that automatically fills in most arguments.

                                  See Fintype.prod_equiv for the version where s and t are univ.

                                  theorem Finset.sum_equiv {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [AddCommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (e : ι κ) (hst : ∀ (i : ι), i s e i t) (hfg : is, f i = g (e i)) :
                                  is, f i = it, g i

                                  Specialization of Finset.sum_nbij'` that automatically fills in most arguments.

                                  See Fintype.sum_equiv for the version where s and t are univ.

                                  theorem Finset.prod_bijective {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [CommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (e : ικ) (he : Function.Bijective e) (hst : ∀ (i : ι), i s e i t) (hfg : is, f i = g (e i)) :
                                  is, f i = it, g i

                                  Specialization of Finset.prod_bij that automatically fills in most arguments.

                                  See Fintype.prod_bijective for the version where s and t are univ.

                                  theorem Finset.sum_bijective {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [AddCommMonoid α] {s : Finset ι} {t : Finset κ} {f : ια} {g : κα} (e : ικ) (he : Function.Bijective e) (hst : ∀ (i : ι), i s e i t) (hfg : is, f i = g (e i)) :
                                  is, f i = it, g i

                                  Specialization of Finset.sum_bij` that automatically fills in most arguments.

                                  See Fintype.sum_bijective for the version where s and t are univ.

                                  theorem Finset.prod_hom_rel {α : Type u_3} {β : Type u_4} {γ : Type u_5} [CommMonoid β] [CommMonoid γ] {r : βγProp} {f : αβ} {g : αγ} {s : Finset α} (h₁ : r 1 1) (h₂ : ∀ (a : α) (b : β) (c : γ), r b cr (f a * b) (g a * c)) :
                                  r (∏ xs, f x) (∏ xs, g x)
                                  theorem Finset.sum_hom_rel {α : Type u_3} {β : Type u_4} {γ : Type u_5} [AddCommMonoid β] [AddCommMonoid γ] {r : βγProp} {f : αβ} {g : αγ} {s : Finset α} (h₁ : r 0 0) (h₂ : ∀ (a : α) (b : β) (c : γ), r b cr (f a + b) (g a + c)) :
                                  r (∑ xs, f x) (∑ xs, g x)
                                  theorem Finset.prod_coe_sort_eq_attach {α : Type u_3} {β : Type u_4} (s : Finset α) [CommMonoid β] (f : { x : α // x s }β) :
                                  i : { x : α // x s }, f i = is.attach, f i
                                  theorem Finset.sum_coe_sort_eq_attach {α : Type u_3} {β : Type u_4} (s : Finset α) [AddCommMonoid β] (f : { x : α // x s }β) :
                                  i : { x : α // x s }, f i = is.attach, f i
                                  theorem Finset.prod_ite_index {α : Type u_3} {β : Type u_4} [CommMonoid β] (p : Prop) [Decidable p] (s t : Finset α) (f : αβ) :
                                  xif p then s else t, f x = if p then xs, f x else xt, f x
                                  theorem Finset.sum_ite_index {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (p : Prop) [Decidable p] (s t : Finset α) (f : αβ) :
                                  xif p then s else t, f x = if p then xs, f x else xt, f x
                                  @[simp]
                                  theorem Finset.prod_ite_irrel {α : Type u_3} {β : Type u_4} [CommMonoid β] (p : Prop) [Decidable p] (s : Finset α) (f g : αβ) :
                                  (∏ xs, if p then f x else g x) = if p then xs, f x else xs, g x
                                  @[simp]
                                  theorem Finset.sum_ite_irrel {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (p : Prop) [Decidable p] (s : Finset α) (f g : αβ) :
                                  (∑ xs, if p then f x else g x) = if p then xs, f x else xs, g x
                                  @[simp]
                                  theorem Finset.prod_dite_irrel {α : Type u_3} {β : Type u_4} [CommMonoid β] (p : Prop) [Decidable p] (s : Finset α) (f : pαβ) (g : ¬pαβ) :
                                  (∏ xs, if h : p then f h x else g h x) = if h : p then xs, f h x else xs, g h x
                                  @[simp]
                                  theorem Finset.sum_dite_irrel {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (p : Prop) [Decidable p] (s : Finset α) (f : pαβ) (g : ¬pαβ) :
                                  (∑ xs, if h : p then f h x else g h x) = if h : p then xs, f h x else xs, g h x
                                  theorem Finset.nonempty_of_prod_ne_one {α : Type u_3} {β : Type u_4} {s : Finset α} {f : αβ} [CommMonoid β] (h : xs, f x 1) :
                                  s.Nonempty
                                  theorem Finset.nonempty_of_sum_ne_zero {α : Type u_3} {β : Type u_4} {s : Finset α} {f : αβ} [AddCommMonoid β] (h : xs, f x 0) :
                                  s.Nonempty
                                  theorem Finset.prod_range_zero {β : Type u_4} [CommMonoid β] (f : β) :
                                  kFinset.range 0, f k = 1
                                  theorem Finset.sum_range_zero {β : Type u_4} [AddCommMonoid β] (f : β) :
                                  kFinset.range 0, f k = 0
                                  theorem Finset.sum_filter_count_eq_countP {α : Type u_3} [DecidableEq α] (p : αProp) [DecidablePred p] (l : List α) :
                                  xFinset.filter (fun (x : α) => p x) l.toFinset, List.count x l = List.countP (fun (b : α) => decide (p b)) l
                                  theorem Finset.prod_mem_multiset {α : Type u_3} {β : Type u_4} [CommMonoid β] [DecidableEq α] (m : Multiset α) (f : { x : α // x m }β) (g : αβ) (hfg : ∀ (x : { x : α // x m }), f x = g x) :
                                  x : { x : α // x m }, f x = xm.toFinset, g x
                                  theorem Finset.sum_mem_multiset {α : Type u_3} {β : Type u_4} [AddCommMonoid β] [DecidableEq α] (m : Multiset α) (f : { x : α // x m }β) (g : αβ) (hfg : ∀ (x : { x : α // x m }), f x = g x) :
                                  x : { x : α // x m }, f x = xm.toFinset, g x
                                  theorem Finset.prod_induction {α : Type u_3} {s : Finset α} {M : Type u_6} [CommMonoid M] (f : αM) (p : MProp) (hom : ∀ (a b : M), p ap bp (a * b)) (unit : p 1) (base : xs, p (f x)) :
                                  p (∏ xs, f x)

                                  To prove a property of a product, it suffices to prove that the property is multiplicative and holds on factors.

                                  theorem Finset.sum_induction {α : Type u_3} {s : Finset α} {M : Type u_6} [AddCommMonoid M] (f : αM) (p : MProp) (hom : ∀ (a b : M), p ap bp (a + b)) (unit : p 0) (base : xs, p (f x)) :
                                  p (∑ xs, f x)

                                  To prove a property of a sum, it suffices to prove that the property is additive and holds on summands.

                                  theorem Finset.prod_induction_nonempty {α : Type u_3} {s : Finset α} {M : Type u_6} [CommMonoid M] (f : αM) (p : MProp) (hom : ∀ (a b : M), p ap bp (a * b)) (nonempty : s.Nonempty) (base : xs, p (f x)) :
                                  p (∏ xs, f x)

                                  To prove a property of a product, it suffices to prove that the property is multiplicative and holds on factors.

                                  theorem Finset.sum_induction_nonempty {α : Type u_3} {s : Finset α} {M : Type u_6} [AddCommMonoid M] (f : αM) (p : MProp) (hom : ∀ (a b : M), p ap bp (a + b)) (nonempty : s.Nonempty) (base : xs, p (f x)) :
                                  p (∑ xs, f x)

                                  To prove a property of a sum, it suffices to prove that the property is additive and holds on summands.

                                  theorem Finset.prod_pow {α : Type u_3} {β : Type u_4} [CommMonoid β] (s : Finset α) (n : ) (f : αβ) :
                                  xs, f x ^ n = (∏ xs, f x) ^ n
                                  theorem Finset.sum_nsmul {α : Type u_3} {β : Type u_4} [AddCommMonoid β] (s : Finset α) (n : ) (f : αβ) :
                                  xs, n f x = n xs, f x
                                  theorem Finset.prod_dvd_prod_of_subset {ι : Type u_6} {M : Type u_7} [CommMonoid M] (s t : Finset ι) (f : ιM) (h : s t) :
                                  is, f i it, f i
                                  @[simp]
                                  theorem Finset.op_sum {α : Type u_3} {β : Type u_4} [AddCommMonoid β] {s : Finset α} (f : αβ) :
                                  MulOpposite.op (∑ xs, f x) = xs, MulOpposite.op (f x)

                                  Moving to the opposite additive commutative monoid commutes with summing.

                                  @[simp]
                                  theorem Finset.unop_sum {α : Type u_3} {β : Type u_4} [AddCommMonoid β] {s : Finset α} (f : αβᵐᵒᵖ) :
                                  MulOpposite.unop (∑ xs, f x) = xs, MulOpposite.unop (f x)
                                  @[simp]
                                  theorem Finset.prod_inv_distrib {α : Type u_3} {β : Type u_4} {s : Finset α} {f : αβ} [DivisionCommMonoid β] :
                                  xs, (f x)⁻¹ = (∏ xs, f x)⁻¹
                                  @[simp]
                                  theorem Finset.sum_neg_distrib {α : Type u_3} {β : Type u_4} {s : Finset α} {f : αβ} [SubtractionCommMonoid β] :
                                  xs, -f x = -xs, f x
                                  @[simp]
                                  theorem Finset.prod_div_distrib {α : Type u_3} {β : Type u_4} {s : Finset α} {f g : αβ} [DivisionCommMonoid β] :
                                  xs, f x / g x = (∏ xs, f x) / xs, g x
                                  @[simp]
                                  theorem Finset.sum_sub_distrib {α : Type u_3} {β : Type u_4} {s : Finset α} {f g : αβ} [SubtractionCommMonoid β] :
                                  xs, (f x - g x) = xs, f x - xs, g x
                                  theorem Finset.prod_zpow {α : Type u_3} {β : Type u_4} [DivisionCommMonoid β] (f : αβ) (s : Finset α) (n : ) :
                                  as, f a ^ n = (∏ as, f a) ^ n
                                  theorem Finset.sum_zsmul {α : Type u_3} {β : Type u_4} [SubtractionCommMonoid β] (f : αβ) (s : Finset α) (n : ) :
                                  as, n f a = n as, f a
                                  theorem Finset.sum_nat_mod {α : Type u_3} (s : Finset α) (n : ) (f : α) :
                                  (∑ is, f i) % n = (∑ is, f i % n) % n
                                  theorem Finset.prod_nat_mod {α : Type u_3} (s : Finset α) (n : ) (f : α) :
                                  (∏ is, f i) % n = (∏ is, f i % n) % n
                                  theorem Finset.sum_int_mod {α : Type u_3} (s : Finset α) (n : ) (f : α) :
                                  (∑ is, f i) % n = (∑ is, f i % n) % n
                                  theorem Finset.prod_int_mod {α : Type u_3} (s : Finset α) (n : ) (f : α) :
                                  (∏ is, f i) % n = (∏ is, f i % n) % n
                                  theorem Fintype.prod_bijective {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [CommMonoid α] (e : ικ) (he : Function.Bijective e) (f : ια) (g : κα) (h : ∀ (x : ι), f x = g (e x)) :
                                  x : ι, f x = x : κ, g x

                                  Fintype.prod_bijective is a variant of Finset.prod_bij that accepts Function.Bijective.

                                  See Function.Bijective.prod_comp for a version without h.

                                  theorem Fintype.sum_bijective {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [AddCommMonoid α] (e : ικ) (he : Function.Bijective e) (f : ια) (g : κα) (h : ∀ (x : ι), f x = g (e x)) :
                                  x : ι, f x = x : κ, g x

                                  Fintype.sum_bijective is a variant of Finset.sum_bij that accepts Function.Bijective.

                                  See Function.Bijective.sum_comp for a version without h.

                                  theorem Function.Bijective.finset_prod {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [CommMonoid α] (e : ικ) (he : Function.Bijective e) (f : ια) (g : κα) (h : ∀ (x : ι), f x = g (e x)) :
                                  x : ι, f x = x : κ, g x

                                  Alias of Fintype.prod_bijective.


                                  Fintype.prod_bijective is a variant of Finset.prod_bij that accepts Function.Bijective.

                                  See Function.Bijective.prod_comp for a version without h.

                                  theorem Function.Bijective.finset_sum {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [AddCommMonoid α] (e : ικ) (he : Function.Bijective e) (f : ια) (g : κα) (h : ∀ (x : ι), f x = g (e x)) :
                                  x : ι, f x = x : κ, g x
                                  theorem Fintype.prod_equiv {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [CommMonoid α] (e : ι κ) (f : ια) (g : κα) (h : ∀ (x : ι), f x = g (e x)) :
                                  x : ι, f x = x : κ, g x

                                  Fintype.prod_equiv is a specialization of Finset.prod_bij that automatically fills in most arguments.

                                  See Equiv.prod_comp for a version without h.

                                  theorem Fintype.sum_equiv {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [AddCommMonoid α] (e : ι κ) (f : ια) (g : κα) (h : ∀ (x : ι), f x = g (e x)) :
                                  x : ι, f x = x : κ, g x

                                  Fintype.sum_equiv is a specialization of Finset.sum_bij that automatically fills in most arguments.

                                  See Equiv.sum_comp for a version without h.

                                  theorem Function.Bijective.prod_comp {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [CommMonoid α] {e : ικ} (he : Function.Bijective e) (g : κα) :
                                  i : ι, g (e i) = i : κ, g i
                                  theorem Function.Bijective.sum_comp {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [AddCommMonoid α] {e : ικ} (he : Function.Bijective e) (g : κα) :
                                  i : ι, g (e i) = i : κ, g i
                                  theorem Equiv.prod_comp {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [CommMonoid α] (e : ι κ) (g : κα) :
                                  i : ι, g (e i) = i : κ, g i
                                  theorem Equiv.sum_comp {ι : Type u_6} {κ : Type u_7} {α : Type u_8} [Fintype ι] [Fintype κ] [AddCommMonoid α] (e : ι κ) (g : κα) :
                                  i : ι, g (e i) = i : κ, g i
                                  theorem Fintype.prod_empty {α : Type u_9} {β : Type u_10} [CommMonoid β] [IsEmpty α] [Fintype α] (f : αβ) :
                                  x : α, f x = 1
                                  theorem Fintype.sum_empty {α : Type u_9} {β : Type u_10} [AddCommMonoid β] [IsEmpty α] [Fintype α] (f : αβ) :
                                  x : α, f x = 0
                                  @[simp]
                                  theorem Finset.prod_attach_univ {ι : Type u_1} {α : Type u_3} [CommMonoid α] [Fintype ι] (f : { i : ι // i Finset.univ }α) :
                                  iFinset.univ.attach, f i = i : ι, f i,
                                  @[simp]
                                  theorem Finset.sum_attach_univ {ι : Type u_1} {α : Type u_3} [AddCommMonoid α] [Fintype ι] (f : { i : ι // i Finset.univ }α) :
                                  iFinset.univ.attach, f i = i : ι, f i,
                                  theorem Finset.prod_erase_attach {ι : Type u_1} {α : Type u_3} [CommMonoid α] [DecidableEq ι] {s : Finset ι} (f : ια) (i : { x : ι // x s }) :
                                  js.attach.erase i, f j = js.erase i, f j
                                  theorem Finset.sum_erase_attach {ι : Type u_1} {α : Type u_3} [AddCommMonoid α] [DecidableEq ι] {s : Finset ι} (f : ια) (i : { x : ι // x s }) :
                                  js.attach.erase i, f j = js.erase i, f j
                                  @[simp]
                                  theorem Multiset.card_sum {ι : Type u_1} {α : Type u_3} (s : Finset ι) (f : ιMultiset α) :
                                  (∑ is, f i).card = is, (f i).card
                                  theorem Multiset.disjoint_list_sum_left {α : Type u_3} {a : Multiset α} {l : List (Multiset α)} :
                                  Disjoint l.sum a bl, Disjoint b a
                                  theorem Multiset.disjoint_list_sum_right {α : Type u_3} {a : Multiset α} {l : List (Multiset α)} :
                                  Disjoint a l.sum bl, Disjoint a b
                                  theorem Multiset.disjoint_sum_left {α : Type u_3} {a : Multiset α} {i : Multiset (Multiset α)} :
                                  Disjoint i.sum a bi, Disjoint b a
                                  theorem Multiset.disjoint_sum_right {α : Type u_3} {a : Multiset α} {i : Multiset (Multiset α)} :
                                  Disjoint a i.sum bi, Disjoint a b
                                  theorem Multiset.disjoint_finset_sum_left {α : Type u_3} {β : Type u_6} {i : Finset β} {f : βMultiset α} {a : Multiset α} :
                                  Disjoint (i.sum f) a bi, Disjoint (f b) a
                                  theorem Multiset.disjoint_finset_sum_right {α : Type u_3} {β : Type u_6} {i : Finset β} {f : βMultiset α} {a : Multiset α} :
                                  Disjoint a (i.sum f) bi, Disjoint a (f b)
                                  theorem Multiset.count_sum' {α : Type u_3} {β : Type u_4} [DecidableEq α] {s : Finset β} {a : α} {f : βMultiset α} :
                                  Multiset.count a (∑ xs, f x) = xs, Multiset.count a (f x)
                                  theorem Multiset.toFinset_prod_dvd_prod {α : Type u_3} [DecidableEq α] [CommMonoid α] (S : Multiset α) :
                                  S.toFinset.prod id S.prod
                                  @[simp]
                                  theorem Units.coe_prod {α : Type u_3} {M : Type u_6} [CommMonoid M] (f : αMˣ) (s : Finset α) :
                                  (∏ is, f i) = is, (f i)

                                  Additive, Multiplicative #

                                  @[simp]
                                  theorem ofMul_list_prod {α : Type u_3} [Monoid α] (s : List α) :
                                  Additive.ofMul s.prod = (List.map (⇑Additive.ofMul) s).sum
                                  @[simp]
                                  theorem toMul_list_sum {α : Type u_3} [Monoid α] (s : List (Additive α)) :
                                  Additive.toMul s.sum = (List.map (⇑Additive.toMul) s).prod
                                  @[simp]
                                  theorem ofAdd_list_prod {α : Type u_3} [AddMonoid α] (s : List α) :
                                  Multiplicative.ofAdd s.sum = (List.map (⇑Multiplicative.ofAdd) s).prod
                                  @[simp]
                                  theorem toAdd_list_sum {α : Type u_3} [AddMonoid α] (s : List (Multiplicative α)) :
                                  Multiplicative.toAdd s.prod = (List.map (⇑Multiplicative.toAdd) s).sum
                                  @[simp]
                                  theorem ofMul_multiset_prod {α : Type u_3} [CommMonoid α] (s : Multiset α) :
                                  Additive.ofMul s.prod = (Multiset.map (⇑Additive.ofMul) s).sum
                                  @[simp]
                                  theorem toMul_multiset_sum {α : Type u_3} [CommMonoid α] (s : Multiset (Additive α)) :
                                  Additive.toMul s.sum = (Multiset.map (⇑Additive.toMul) s).prod
                                  @[simp]
                                  theorem ofMul_prod {ι : Type u_1} {α : Type u_3} [CommMonoid α] (s : Finset ι) (f : ια) :
                                  Additive.ofMul (∏ is, f i) = is, Additive.ofMul (f i)
                                  @[simp]
                                  theorem toMul_sum {ι : Type u_1} {α : Type u_3} [CommMonoid α] (s : Finset ι) (f : ιAdditive α) :
                                  Additive.toMul (∑ is, f i) = is, Additive.toMul (f i)
                                  @[simp]
                                  theorem ofAdd_multiset_prod {α : Type u_3} [AddCommMonoid α] (s : Multiset α) :
                                  Multiplicative.ofAdd s.sum = (Multiset.map (⇑Multiplicative.ofAdd) s).prod
                                  @[simp]
                                  theorem toAdd_multiset_sum {α : Type u_3} [AddCommMonoid α] (s : Multiset (Multiplicative α)) :
                                  Multiplicative.toAdd s.prod = (Multiset.map (⇑Multiplicative.toAdd) s).sum
                                  @[simp]
                                  theorem ofAdd_sum {ι : Type u_1} {α : Type u_3} [AddCommMonoid α] (s : Finset ι) (f : ια) :
                                  Multiplicative.ofAdd (∑ is, f i) = is, Multiplicative.ofAdd (f i)
                                  @[simp]
                                  theorem toAdd_prod {ι : Type u_1} {α : Type u_3} [AddCommMonoid α] (s : Finset ι) (f : ιMultiplicative α) :
                                  Multiplicative.toAdd (∏ is, f i) = is, Multiplicative.toAdd (f i)