Affine combinations of points #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file defines affine combinations of points.
Main definitions #
-
weighted_vsub_of_point
is a general weighted combination of subtractions with an explicit base point, yielding a vector. -
weighted_vsub
uses an arbitrary choice of base point and is intended to be used when the sum of weights is 0, in which case the result is independent of the choice of base point. -
affine_combination
adds the weighted combination to the arbitrary base point, yielding a point rather than a vector, and is intended to be used when the sum of weights is 1, in which case the result is independent of the choice of base point.
These definitions are for sums over a finset
; versions for a
fintype
may be obtained using finset.univ
, while versions for a
finsupp
may be obtained using finsupp.support
.
References #
A weighted sum of the results of subtracting a base point from the given points, as a linear map on the weights. The main cases of interest are where the sum of the weights is 0, in which case the sum is independent of the choice of base point, and where the sum of the weights is 1, in which case the sum added to the base point is independent of the choice of base point.
Equations
- s.weighted_vsub_of_point p b = s.sum (λ (i : ι), (linear_map.proj i).smul_right (p i -ᵥ b))
The value of weighted_vsub_of_point
, where the given points are equal.
weighted_vsub_of_point
gives equal results for two families of weights and two families of
points that are equal on s
.
Given a family of points, if we use a member of the family as a base point, the
weighted_vsub_of_point
does not depend on the value of the weights at this point.
The weighted sum is independent of the base point when the sum of the weights is 0.
The weighted sum, added to the base point, is independent of the base point when the sum of the weights is 1.
The weighted sum is unaffected by removing the base point, if present, from the set of points.
The weighted sum is unaffected by adding the base point, whether or not present, to the set of points.
The weighted sum is unaffected by changing the weights to the corresponding indicator function and adding points to the set.
A weighted sum, over the image of an embedding, equals a weighted
sum with the same points and weights over the original
finset
.
A weighted sum of pairwise subtractions, expressed as a subtraction of two
weighted_vsub_of_point
expressions.
A weighted sum of pairwise subtractions, where the point on the right is constant,
expressed as a subtraction involving a weighted_vsub_of_point
expression.
A weighted sum of pairwise subtractions, where the point on the left is constant,
expressed as a subtraction involving a weighted_vsub_of_point
expression.
A weighted sum may be split into such sums over two subsets.
A weighted sum may be split into a subtraction of such sums over two subsets.
A weighted sum over s.subtype pred
equals one over s.filter pred
.
A weighted sum over s.filter pred
equals one over s
if all the weights at indices in s
not satisfying pred
are zero.
A constant multiplier of the weights in weighted_vsub_of_point
may be moved outside the
sum.
A weighted sum of the results of subtracting a default base point from the given points, as a linear map on the weights. This is intended to be used when the sum of the weights is 0; that condition is specified as a hypothesis on those lemmas that require it.
Equations
- s.weighted_vsub p = s.weighted_vsub_of_point p (classical.choice finset.weighted_vsub._proof_1)
Applying weighted_vsub
with given weights. This is for the case
where a result involving a default base point is OK (for example, when
that base point will cancel out later); a more typical use case for
weighted_vsub
would involve selecting a preferred base point with
weighted_vsub_eq_weighted_vsub_of_point_of_sum_eq_zero
and then
using weighted_vsub_of_point_apply
.
weighted_vsub
gives the sum of the results of subtracting any
base point, when the sum of the weights is 0.
The value of weighted_vsub
, where the given points are equal and the sum of the weights
is 0.
The weighted_vsub
for an empty set is 0.
weighted_vsub
gives equal results for two families of weights and two families of points
that are equal on s
.
The weighted sum is unaffected by changing the weights to the corresponding indicator function and adding points to the set.
A weighted subtraction, over the image of an embedding, equals a
weighted subtraction with the same points and weights over the
original finset
.
A weighted sum of pairwise subtractions, expressed as a subtraction of two weighted_vsub
expressions.
A weighted sum of pairwise subtractions, where the point on the right is constant and the sum of the weights is 0.
A weighted sum of pairwise subtractions, where the point on the left is constant and the sum of the weights is 0.
A weighted sum may be split into such sums over two subsets.
A weighted sum may be split into a subtraction of such sums over two subsets.
A weighted sum over s.subtype pred
equals one over s.filter pred
.
A weighted sum over s.filter pred
equals one over s
if all the weights at indices in s
not satisfying pred
are zero.
A constant multiplier of the weights in weighted_vsub_of
may be moved outside the sum.
A weighted sum of the results of subtracting a default base point from the given points, added to that base point, as an affine map on the weights. This is intended to be used when the sum of the weights is 1, in which case it is an affine combination (barycenter) of the points with the given weights; that condition is specified as a hypothesis on those lemmas that require it.
Equations
- finset.affine_combination k s p = {to_fun := λ (w : ι → k), ⇑(s.weighted_vsub_of_point p (classical.choice finset.affine_combination._proof_1)) w +ᵥ classical.choice finset.affine_combination._proof_1, linear := s.weighted_vsub p, map_vadd' := _}
The linear map corresponding to affine_combination
is
weighted_vsub
.
Applying affine_combination
with given weights. This is for the
case where a result involving a default base point is OK (for example,
when that base point will cancel out later); a more typical use case
for affine_combination
would involve selecting a preferred base
point with
affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one
and
then using weighted_vsub_of_point_apply
.
The value of affine_combination
, where the given points are equal.
affine_combination
gives equal results for two families of weights and two families of
points that are equal on s
.
affine_combination
gives the sum with any base point, when the
sum of the weights is 1.
Adding a weighted_vsub
to an affine_combination
.
Subtracting two affine_combination
s.
Viewing a module as an affine space modelled on itself, a weighted_vsub
is just a linear
combination.
Viewing a module as an affine space modelled on itself, affine combinations are just linear combinations.
An affine_combination
equals a point if that point is in the set
and has weight 1 and the other points in the set have weight 0.
An affine combination is unaffected by changing the weights to the corresponding indicator function and adding points to the set.
An affine combination, over the image of an embedding, equals an
affine combination with the same points and weights over the original
finset
.
A weighted sum of pairwise subtractions, expressed as a subtraction of two affine_combination
expressions.
A weighted sum of pairwise subtractions, where the point on the right is constant and the sum of the weights is 1.
A weighted sum of pairwise subtractions, where the point on the left is constant and the sum of the weights is 1.
A weighted sum may be split into a subtraction of affine combinations over two subsets.
If a weighted sum is zero and one of the weights is -1
, the corresponding point is
the affine combination of the other points with the given weights.
An affine combination over s.subtype pred
equals one over s.filter pred
.
An affine combination over s.filter pred
equals one over s
if all the weights at indices
in s
not satisfying pred
are zero.
Suppose an indexed family of points is given, along with a subset
of the index type. A vector can be expressed as
weighted_vsub_of_point
using a finset
lying within that subset and
with a given sum of weights if and only if it can be expressed as
weighted_vsub_of_point
with that sum of weights for the
corresponding indexed family whose index type is the subtype
corresponding to that subset.
Suppose an indexed family of points is given, along with a subset
of the index type. A vector can be expressed as weighted_vsub
using
a finset
lying within that subset and with sum of weights 0 if and
only if it can be expressed as weighted_vsub
with sum of weights 0
for the corresponding indexed family whose index type is the subtype
corresponding to that subset.
Suppose an indexed family of points is given, along with a subset
of the index type. A point can be expressed as an
affine_combination
using a finset
lying within that subset and
with sum of weights 1 if and only if it can be expressed an
affine_combination
with sum of weights 1 for the corresponding
indexed family whose index type is the subtype corresponding to that
subset.
Affine maps commute with affine combinations.
Weights for expressing a single point as an affine combination.
Equations
Weights for expressing the subtraction of two points as a weighted_vsub
.
Weights for expressing line_map
as an affine combination.
Equations
An affine combination with affine_combination_single_weights
gives the specified point.
A weighted subtraction with weighted_vsub_vsub_weights
gives the result of subtracting the
specified points.
An affine combination with affine_combination_line_map_weights
gives the result of
line_map
.
The weights for the centroid of some points.
Equations
- finset.centroid_weights k s = function.const ι (↑(s.card))⁻¹
centroid_weights
at any point.
centroid_weights
equals a constant function.
The weights in the centroid sum to 1, if the number of points,
converted to k
, is not zero.
In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is not zero.
In the characteristic zero case, the weights in the centroid sum to 1 if the set is nonempty.
In the characteristic zero case, the weights in the centroid sum
to 1 if the number of points is n + 1
.
The centroid of some points. Although defined for any s
, this
is intended to be used in the case where the number of points,
converted to k
, is not zero.
Equations
- finset.centroid k s p = ⇑(finset.affine_combination k s p) (finset.centroid_weights k s)
The definition of the centroid.
The centroid of a single point.
The centroid of two points, expressed directly as adding a vector to a point.
The centroid of two points indexed by fin 2
, expressed directly
as adding a vector to the first point.
A centroid, over the image of an embedding, equals a centroid with
the same points and weights over the original finset
.
centroid_weights
gives the weights for the centroid as a
constant function, which is suitable when summing over the points
whose centroid is being taken. This function gives the weights in a
form suitable for summing over a larger set of points, as an indicator
function that is zero outside the set whose centroid is being taken.
In the case of a fintype
, the sum may be over univ
.
Equations
The definition of centroid_weights_indicator
.
The sum of the weights for the centroid indexed by a fintype
.
In the characteristic zero case, the weights in the centroid
indexed by a fintype
sum to 1 if the number of points is not
zero.
In the characteristic zero case, the weights in the centroid
indexed by a fintype
sum to 1 if the set is nonempty.
In the characteristic zero case, the weights in the centroid
indexed by a fintype
sum to 1 if the number of points is n + 1
.
The centroid as an affine combination over a fintype
.
An indexed family of points that is injective on the given
finset
has the same centroid as the image of that finset
. This is
stated in terms of a set equal to the image to provide control of
definitional equality for the index type used for the centroid of the
image.
Two indexed families of points that are injective on the given
finset
s and with the same points in the image of those finset
s
have the same centroid.
A weighted_vsub
with sum of weights 0 is in the vector_span
of
an indexed family.
An affine_combination
with sum of weights 1 is in the
affine_span
of an indexed family, if the underlying ring is
nontrivial.
A vector is in the vector_span
of an indexed family if and only
if it is a weighted_vsub
with sum of weights 0.
A point in the affine_span
of an indexed family is an
affine_combination
with sum of weights 1. See also
eq_affine_combination_of_mem_affine_span_of_fintype
.
A point is in the affine_span
of an indexed family if and only
if it is an affine_combination
with sum of weights 1, provided the
underlying ring is nontrivial.
Given a family of points together with a chosen base point in that family, membership of the
affine span of this family corresponds to an identity in terms of weighted_vsub_of_point
, with
weights that are not required to sum to 1.
Given a set of points, together with a chosen base point in this set, if we affinely transport all other members of the set along the line joining them to this base point, the affine span is unchanged.
The centroid lies in the affine span if the number of points,
converted to k
, is not zero.
In the characteristic zero case, the centroid lies in the affine span if the number of points is not zero.
In the characteristic zero case, the centroid lies in the affine span if the set is nonempty.
In the characteristic zero case, the centroid lies in the affine
span if the number of points is n + 1
.
A weighted sum, as an affine map on the points involved.
Equations
- affine_map.weighted_vsub_of_point P s w = {to_fun := λ (p : (ι → P) × P), ⇑(s.weighted_vsub_of_point p.fst p.snd) w, linear := s.sum (λ (i : ι), w i • ((linear_map.proj i).comp (linear_map.fst k (ι → V) V) - linear_map.snd k (ι → V) V)), map_vadd' := _}