Affine combinations of points #
This file defines affine combinations of points.
Main definitions #
weightedVSubOfPointis a general weighted combination of subtractions with an explicit base point, yielding a vector.weightedVSubuses 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.affineCombinationadds 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.weightedVSubOfPoint p b = ∑ i ∈ s, (LinearMap.proj i).smulRight (p i -ᵥ b)
Instances For
The value of weightedVSubOfPoint, where the given points are equal.
weightedVSubOfPoint 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 weightedVSubOfPoint 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 weightedVSubOfPoint expressions.
A weighted sum of pairwise subtractions, where the point on the right is constant, expressed as a subtraction involving a weightedVSubOfPoint expression.
A weighted sum of pairwise subtractions, where the point on the left is constant, expressed as a subtraction involving a weightedVSubOfPoint 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 {x ∈ s | pred x}.
A weighted sum over {x ∈ s | pred x} equals one over s if all the weights at indices in s not satisfying pred are zero.
A constant multiplier of the weights in weightedVSubOfPoint 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.weightedVSub p = s.weightedVSubOfPoint p (Classical.choice ⋯)
Instances For
Applying weightedVSub 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 weightedVSub would involve selecting a preferred base point with weightedVSub_eq_weightedVSubOfPoint_of_sum_eq_zero and then using weightedVSubOfPoint_apply.
weightedVSub gives the sum of the results of subtracting any base point, when the sum of the weights is 0.
The value of weightedVSub, where the given points are equal and the sum of the weights is 0.
The weightedVSub for an empty set is 0.
weightedVSub 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 weightedVSub 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 {x ∈ s | pred x}.
A weighted sum over {x ∈ s | pred x} equals one over s if all the weights at indices in s not satisfying pred are zero.
A constant multiplier of the weights in weightedVSub_of may be moved outside the sum.
Equations
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.affineCombination k s p = { toFun := fun (w : ι → k) => (s.weightedVSubOfPoint p (Classical.choice ⋯)) w +ᵥ Classical.choice ⋯, linear := s.weightedVSub p, map_vadd' := ⋯ }
Instances For
The linear map corresponding to affineCombination is weightedVSub.
Applying affineCombination 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 affineCombination would involve selecting a preferred base point with affineCombination_eq_weightedVSubOfPoint_vadd_of_sum_eq_one and then using weightedVSubOfPoint_apply.
The value of affineCombination, where the given points are equal.
affineCombination gives equal results for two families of weights and two families of points that are equal on s.
affineCombination gives the sum with any base point, when the sum of the weights is 1.
Adding a weightedVSub to an affineCombination.
Subtracting two affineCombinations.
Viewing a module as an affine space modelled on itself, a weightedVSub is just a linear combination.
Viewing a module as an affine space modelled on itself, affine combinations are just linear combinations.
An affineCombination 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 affineCombination 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 {x ∈ s | pred x}.
An affine combination over {x ∈ s | pred x} 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 weightedVSubOfPoint using a Finset lying within that subset and with a given sum of weights if and only if it can be expressed as weightedVSubOfPoint 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 weightedVSub using a Finset lying within that subset and with sum of weights 0 if and only if it can be expressed as weightedVSub 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 affineCombination using a Finset lying within that subset and with sum of weights 1 if and only if it can be expressed an affineCombination 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.
The value of affineCombination, where the given points take only two values.
Weights for expressing a single point as an affine combination.
Equations
Instances For
Weights for expressing the subtraction of two points as a weightedVSub.
Equations
Instances For
Weights for expressing lineMap as an affine combination.
Equations
Instances For
An affine combination with affineCombinationSingleWeights gives the specified point.
A weighted subtraction with weightedVSubVSubWeights gives the result of subtracting the specified points.
An affine combination with affineCombinationLineMapWeights gives the result of line_map.
A weightedVSub with sum of weights 0 is in the vectorSpan of an indexed family.
An affineCombination with sum of weights 1 is in the affineSpan of an indexed family, if the underlying ring is nontrivial.
An affineCombination with sum of weights 1 is in the affineSpan of an indexed family, if the family is nonempty.
A vector is in the vectorSpan of an indexed family if and only if it is a weightedVSub with sum of weights 0.
A point in the affineSpan of an indexed family is an affineCombination with sum of weights 1. See also eq_affineCombination_of_mem_affineSpan_of_fintype.
A point in the affineSpan of a subset of an indexed family is an affineCombination with sum of weights 1, using only points in the given subset.
A point is in the affineSpan of an indexed family if and only if it is an affineCombination 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 weightedVSubOfPoint, 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.
A weighted sum, as an affine map on the points involved.
Equations
- One or more equations did not get rendered due to their size.