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 x
is notation forFinset.prod s f
(assumingβ
is aCommMonoid
)∑ x ∈ s, f x
is notation forFinset.sum s f
(assumingβ
is anAddCommMonoid
)∏ x, f x
is notation forFinset.prod Finset.univ f
(assumingα
is aFintype
andβ
is aCommMonoid
)∑ x, f x
is notation forFinset.sum Finset.univ f
(assumingα
is aFintype
andβ
is anAddCommMonoid
)
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 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
- s.prod f = (Multiset.map f s.val).prod
Instances For
∑ 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
- s.sum f = (Multiset.map f s.val).sum
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 x
is notation forFinset.sum Finset.univ f
. It is the sum off x
, wherex
ranges over the finite domain off
.∑ x ∈ s, f x
is notation forFinset.sum s f
. It is the sum off x
, wherex
ranges over the finite sets
(either aFinset
or aSet
with aFintype
instance).∑ x ∈ s with p x, f x
is notation forFinset.sum (Finset.filter p s) f
.∑ (x ∈ s) (y ∈ t), f x y
is notation forFinset.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 forFinset.prod Finset.univ f
. It is the product off x
, wherex
ranges over the finite domain off
.∏ x ∈ s, f x
is notation forFinset.prod s f
. It is the product off x
, wherex
ranges over the finite sets
(either aFinset
or aSet
with aFintype
instance).∏ x ∈ s with p x, f x
is notation forFinset.prod (Finset.filter p s) f
.∏ (x ∈ s) (y ∈ t), f x y
is notation forFinset.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
Alias of Finset.prod_of_isEmpty
.
Alias of Finset.sum_of_isEmpty
.
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 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 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_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_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 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 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.
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.
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_nbij'` that automatically fills in most arguments.
See Fintype.sum_equiv
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
.
Specialization of
Finset.sum_bij` that automatically fills in most arguments.
See Fintype.sum_bijective
for the version where s
and t
are univ
.
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.
To prove a property of a sum, it suffices to prove that the property is additive and holds on summands.
Moving to the opposite additive commutative monoid commutes with summing.
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_bijective
is a variant of Finset.sum_bij
that accepts
Function.Bijective
.
See Function.Bijective.sum_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.prod_equiv
is a specialization of Finset.prod_bij
that
automatically fills in most arguments.
See Equiv.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
.