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.
- ∏ x ∈ s, f xis notation for- Finset.prod s f(assuming- βis a- CommMonoid)
- ∑ x ∈ s, f xis notation for- Finset.sum s f(assuming- βis an- AddCommMonoid)
- ∏ x, f xis notation for- Finset.prod Finset.univ f(assuming- αis a- Fintypeand- βis a- CommMonoid)
- ∑ x, f xis notation for- Finset.sum Finset.univ f(assuming- αis a- Fintypeand- βis an- AddCommMonoid)
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.
∑ x ∈ s, f x is the sum of f x as x ranges over the elements
of the finite set s.
Equations
- s.sum f = (Multiset.map f s.val).sum
Instances For
∏ x ∈ s, f x is the product of f x
as x ranges over the elements of the finite set s.
Equations
- s.prod f = (Multiset.map f s.val).prod
Instances For
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
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
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 xis notation for- Finset.sum Finset.univ f. It is the sum of- f x, where- xranges over the finite domain of- f.
- ∑ x ∈ s, f xis notation for- Finset.sum s f. It is the sum of- f x, where- xranges over the finite set- s(either a- Finsetor a- Setwith a- Fintypeinstance).
- ∑ x ∈ s with p x, f xis notation for- Finset.sum (Finset.filter p s) f.
- ∑ (x ∈ s) (y ∈ t), f x yis 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 xis notation for- Finset.prod Finset.univ f. It is the product of- f x, where- xranges over the finite domain of- f.
- ∏ x ∈ s, f xis notation for- Finset.prod s f. It is the product of- f x, where- xranges over the finite set- s(either a- Finsetor a- Setwith a- Fintypeinstance).
- ∏ x ∈ s with p x, f xis notation for- Finset.prod (Finset.filter p s) f.
- ∏ (x ∈ s) (y ∈ t), f x yis 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
See also Finset.sum_apply, with the same conclusion but with the weaker hypothesis
f : α → β → γ
See also Finset.prod_apply, with the same conclusion but with the weaker hypothesis
f : α → β → γ
Alias of Finset.prod_of_isEmpty.
Alias of Finset.sum_of_isEmpty.
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.
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.
The sum of f over insert a s is the same as
the sum over s, as long as f a = 0.
The product of f over insert a s is the same as
the product over s, as long as f a = 1.
A sum over all subsets of s ∪ {x} is obtained by summing the sum over all subsets
of s, and over all subsets of s to which one adds x.
A product over all subsets of s ∪ {x} is obtained by multiplying the product over all subsets
of s, and over all subsets of s to which one adds x.
A sum over all subsets of s ∪ {x} is obtained by summing the sum over all subsets
of s, and over all subsets of s to which one adds x.
A product over all subsets of s ∪ {x} is obtained by multiplying the product over all subsets
of s, and over all subsets of s to which one adds x.
A sum over powerset s is equal to the double sum over sets of subsets of s with
card s = k, for k = 1, ..., card s
A product over powerset s is equal to the double product over sets of subsets of s with
card s = k, for k = 1, ..., card s.
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.
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.
Sum over a sigma type equals the sum of fiberwise sums. For rewriting
in the reverse direction, use Finset.sum_sigma'
Product over a sigma type equals the product of fiberwise products. For rewriting
in the reverse direction, use Finset.prod_sigma'.
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.
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.
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.
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.
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.
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.
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.
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.
Specialization of Finset.sum_nbij'` that automatically fills in most arguments.
See Fintype.sum_equiv for the version where s and t are univ.
Specialization of Finset.prod_nbij' that automatically fills in most arguments.
See Fintype.prod_equiv for the version where s and t are univ.
Specialization of Finset.sum_bij` that automatically fills in most arguments.
See Fintype.sum_bijective for the version where s and t are univ.
Specialization of Finset.prod_bij that automatically fills in most arguments.
See Fintype.prod_bijective for the version where s and t are univ.
Taking a sum over univ.pi t is the same as taking the sum over
Fintype.piFinset t. univ.pi t and Fintype.piFinset t are essentially the same Finset,
but differ in the type of their element, univ.pi t is a Finset (Π a ∈ univ, t a) and
Fintype.piFinset t is a Finset (Π a, t a).
Taking a product over univ.pi t is the same as taking the product over Fintype.piFinset t.
univ.pi t and Fintype.piFinset t are essentially the same Finset, but differ
in the type of their element, univ.pi t is a Finset (Π a ∈ univ, t a) and
Fintype.piFinset t is a Finset (Π a, t a).
An uncurried version of Finset.sum_product
An uncurried version of Finset.prod_product.
An uncurried version of Finset.sum_product_right
An uncurried version of Finset.prod_product_right.
Generalization of Finset.sum_comm to the case when the inner Finsets depend on
the outer variable.
Generalization of Finset.prod_comm to the case when the inner Finsets depend on the outer
variable.
A sum over s.subtype p equals one over s.filter p.
A product over s.subtype p equals one over s.filter p.
If all elements of a Finset satisfy the predicate p, a sum
over s.subtype p equals that sum over s.
If all elements of a Finset satisfy the predicate p, a product
over s.subtype p equals that product over s.
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.
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.
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.
Equations
- ⋯ = ⋯
Instances For
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.
A sum taken over a conditional whose condition is an equality
test on the index and whose alternative is 0 has value either the term at that index or 0.
The difference with Finset.sum_ite_eq is that the arguments to Eq are swapped.
A product taken over a conditional whose condition is an equality test on the index and whose
alternative is 1 has value either the term at that index or 1.
The difference with Finset.prod_ite_eq is that the arguments to Eq are swapped.
Consider a sum of g i (f i) over a finset.  Suppose g is a function such as
n ↦ (n • ·), which maps a second argument of 0 to 0 (or a weighted sum of f i * h i or
f i • h i, where f gives the weights that are multiplied by some other function h). Then if
f is replaced by the corresponding indicator function, the finset may be replaced by a possibly
larger finset without changing the value of the sum.
Consider a product of g i (f i) over a finset.  Suppose g is a function such as
n ↦ (· ^ n), which maps a second argument of 1 to 1. Then if f is replaced by the
corresponding multiplicative indicator function, the finset may be replaced by a possibly larger
finset without changing the value of the product.
Summing an indicator function over a possibly larger Finset is the same as summing
the original function over the original finset.
Taking the product of an indicator function over a possibly larger finset is the same as taking the original function over the original finset.
To prove a property of a sum, it suffices to prove that the property is additive and holds on summands.
To prove a property of a product, it suffices to prove that the property is multiplicative and holds on factors.
To prove a property of a sum, it suffices to prove that the property is additive and holds on summands.
To prove a property of a product, it suffices to prove that the property is multiplicative and holds on factors.
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.
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.
A telescoping sum along {0, ..., n - 1} of an additive commutative group valued
function reduces to the difference of the last and first terms.
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.
A sum over Finset.powersetCard which only depends on the size of the sets is constant.
A product over Finset.powersetCard which only depends on the size of the sets is constant.
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.
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.
A sum can be partitioned into a sum of sums, each equivalent under a setoid.
A product can be partitioned into a product of products, each equivalent under a setoid.
If we can partition a sum into subsets that cancel out, then the whole sum cancels.
If we can partition a product into subsets that cancel out, then the whole product cancels.
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.
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.
A variant of Finset.add_sum_erase with the addition swapped.
A variant of Finset.mul_prod_erase with the multiplication swapped.
If a function applied at a point is 0, a sum is unchanged by
removing that point, if present, from a Finset.
If a function applied at a point is 1, a product is unchanged by
removing that point, if present, from a Finset.
See also Finset.sum_boole.
See also Finset.prod_boole.
If a sum is 0 and the function is 0 except possibly at one
point, it is 0 everywhere on the Finset.
If a product is 1 and the function is 1 except possibly at one
point, it is 1 everywhere on the Finset.
Moving to the opposite additive commutative monoid commutes with summing.
Fintype.sum_bijective is a variant of Finset.sum_bij that accepts
Function.Bijective.
See Function.Bijective.sum_comp for a version without h.
Fintype.prod_bijective is a variant of Finset.prod_bij that accepts Function.Bijective.
See Function.Bijective.prod_comp for a version without h.
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.
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.
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.
See also Finset.sum_dite_eq.
See also Finset.prod_dite_eq.
See also Finset.sum_dite_eq'.
See also Finset.prod_dite_eq'.
See also Finset.sum_ite_eq.
See also Finset.prod_ite_eq.
See also Finset.sum_ite_eq'.
See also Finset.prod_ite_eq'.
See also Finset.sum_pi_single.
See also Finset.prod_pi_mulSingle.
See also Finset.sum_pi_single'.
See also Finset.prod_pi_mulSingle'.
Alias of Fintype.prod_equiv.
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.
Alias of Fintype.sum_equiv.
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.