mathlib documentation

linear_algebra.affine_space.combination

Affine combinations of points #

This file defines affine combinations of points.

Main definitions #

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 #

theorem finset.univ_fin2  :
finset.univ = {0, 1}
def finset.weighted_vsub_of_point {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) (b : P) :
(ι → k) →ₗ[k] V

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
@[simp]
theorem finset.weighted_vsub_of_point_apply {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) (b : P) :
(s.weighted_vsub_of_point p b) w = s.sum (λ (i : ι), w i (p i -ᵥ b))
@[simp]
theorem finset.weighted_vsub_of_point_apply_const {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p b : P) :
(s.weighted_vsub_of_point (λ (_x : ι), p) b) w = s.sum (λ (i : ι), w i) (p -ᵥ b)

The value of weighted_vsub_of_point, where the given points are equal.

theorem finset.weighted_vsub_of_point_eq_of_weights_eq {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) (j : ι) (w₁ w₂ : ι → k) (hw : ∀ (i : ι), i jw₁ i = w₂ i) :
(s.weighted_vsub_of_point p (p j)) w₁ = (s.weighted_vsub_of_point p (p j)) w₂

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.

theorem finset.weighted_vsub_of_point_eq_of_sum_eq_zero {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) (h : s.sum (λ (i : ι), w i) = 0) (b₁ b₂ : P) :

The weighted sum is independent of the base point when the sum of the weights is 0.

theorem finset.weighted_vsub_of_point_vadd_eq_of_sum_eq_one {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) (h : s.sum (λ (i : ι), w i) = 1) (b₁ b₂ : P) :
(s.weighted_vsub_of_point p b₁) w +ᵥ b₁ = (s.weighted_vsub_of_point p b₂) w +ᵥ b₂

The weighted sum, added to the base point, is independent of the base point when the sum of the weights is 1.

@[simp]
theorem finset.weighted_vsub_of_point_erase {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) (i : ι) :

The weighted sum is unaffected by removing the base point, if present, from the set of points.

@[simp]
theorem finset.weighted_vsub_of_point_insert {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) [decidable_eq ι] (w : ι → k) (p : ι → P) (i : ι) :

The weighted sum is unaffected by adding the base point, whether or not present, to the set of points.

theorem finset.weighted_vsub_of_point_indicator_subset {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (w : ι → k) (p : ι → P) (b : P) {s₁ s₂ : finset ι} (h : s₁ s₂) :

The weighted sum is unaffected by changing the weights to the corresponding indicator function and adding points to the set.

theorem finset.weighted_vsub_of_point_map {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} {ι₂ : Type u_5} (s₂ : finset ι₂) (e : ι₂ ι) (w : ι → k) (p : ι → P) (b : P) :

A weighted sum, over the image of an embedding, equals a weighted sum with the same points and weights over the original finset.

theorem finset.sum_smul_vsub_eq_weighted_vsub_of_point_sub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₁ p₂ : ι → P) (b : P) :
s.sum (λ (i : ι), w i (p₁ i -ᵥ p₂ i)) = (s.weighted_vsub_of_point p₁ b) w - (s.weighted_vsub_of_point p₂ b) w

A weighted sum of pairwise subtractions, expressed as a subtraction of two weighted_vsub_of_point expressions.

theorem finset.sum_smul_vsub_const_eq_weighted_vsub_of_point_sub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₁ : ι → P) (p₂ b : P) :
s.sum (λ (i : ι), w i (p₁ i -ᵥ p₂)) = (s.weighted_vsub_of_point p₁ b) w - s.sum (λ (i : ι), w i) (p₂ -ᵥ b)

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.

theorem finset.sum_smul_const_vsub_eq_sub_weighted_vsub_of_point {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₂ : ι → P) (p₁ b : P) :
s.sum (λ (i : ι), w i (p₁ -ᵥ p₂ i)) = s.sum (λ (i : ι), w i) (p₁ -ᵥ b) - (s.weighted_vsub_of_point p₂ b) w

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.

noncomputable def finset.weighted_vsub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) :
(ι → k) →ₗ[k] V

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
theorem finset.weighted_vsub_apply {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) :

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.

theorem finset.weighted_vsub_eq_weighted_vsub_of_point_of_sum_eq_zero {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) (h : s.sum (λ (i : ι), w i) = 0) (b : P) :

weighted_vsub gives the sum of the results of subtracting any base point, when the sum of the weights is 0.

@[simp]
theorem finset.weighted_vsub_apply_const {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : P) (h : s.sum (λ (i : ι), w i) = 0) :
(s.weighted_vsub (λ (_x : ι), p)) w = 0

The value of weighted_vsub, where the given points are equal and the sum of the weights is 0.

@[simp]
theorem finset.weighted_vsub_empty {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (w : ι → k) (p : ι → P) :

The weighted_vsub for an empty set is 0.

theorem finset.weighted_vsub_indicator_subset {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (w : ι → k) (p : ι → P) {s₁ s₂ : finset ι} (h : s₁ s₂) :
(s₁.weighted_vsub p) w = (s₂.weighted_vsub p) (s₁.indicator w)

The weighted sum is unaffected by changing the weights to the corresponding indicator function and adding points to the set.

theorem finset.weighted_vsub_map {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} {ι₂ : Type u_5} (s₂ : finset ι₂) (e : ι₂ ι) (w : ι → k) (p : ι → P) :
((finset.map e s₂).weighted_vsub p) w = (s₂.weighted_vsub (p e)) (w e)

A weighted subtraction, over the image of an embedding, equals a weighted subtraction with the same points and weights over the original finset.

theorem finset.sum_smul_vsub_eq_weighted_vsub_sub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₁ p₂ : ι → P) :
s.sum (λ (i : ι), w i (p₁ i -ᵥ p₂ i)) = (s.weighted_vsub p₁) w - (s.weighted_vsub p₂) w

A weighted sum of pairwise subtractions, expressed as a subtraction of two weighted_vsub expressions.

theorem finset.sum_smul_vsub_const_eq_weighted_vsub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₁ : ι → P) (p₂ : P) (h : s.sum (λ (i : ι), w i) = 0) :
s.sum (λ (i : ι), w i (p₁ i -ᵥ p₂)) = (s.weighted_vsub p₁) w

A weighted sum of pairwise subtractions, where the point on the right is constant and the sum of the weights is 0.

theorem finset.sum_smul_const_vsub_eq_neg_weighted_vsub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₂ : ι → P) (p₁ : P) (h : s.sum (λ (i : ι), w i) = 0) :
s.sum (λ (i : ι), w i (p₁ -ᵥ p₂ i)) = -(s.weighted_vsub p₂) w

A weighted sum of pairwise subtractions, where the point on the left is constant and the sum of the weights is 0.

noncomputable def finset.affine_combination {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) :
(ι → k) →ᵃ[k] P

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
@[simp]
theorem finset.affine_combination_linear {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) :

The linear map corresponding to affine_combination is weighted_vsub.

theorem finset.affine_combination_apply {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) :

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.

@[simp]
theorem finset.affine_combination_apply_const {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : P) (h : s.sum (λ (i : ι), w i) = 1) :
(s.affine_combination (λ (_x : ι), p)) w = p

The value of affine_combination, where the given points are equal.

theorem finset.affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) (h : s.sum (λ (i : ι), w i) = 1) (b : P) :

affine_combination gives the sum with any base point, when the sum of the weights is 1.

theorem finset.weighted_vsub_vadd_affine_combination {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w₁ w₂ : ι → k) (p : ι → P) :
(s.weighted_vsub p) w₁ +ᵥ (s.affine_combination p) w₂ = (s.affine_combination p) (w₁ + w₂)

Adding a weighted_vsub to an affine_combination.

theorem finset.affine_combination_vsub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w₁ w₂ : ι → k) (p : ι → P) :
(s.affine_combination p) w₁ -ᵥ (s.affine_combination p) w₂ = (s.weighted_vsub p) (w₁ - w₂)

Subtracting two affine_combinations.

theorem finset.attach_affine_combination_of_injective {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] (s : finset P) (w : P → k) (f : s → P) (hf : function.injective f) :
theorem finset.attach_affine_combination_coe {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] (s : finset P) (w : P → k) :
@[simp]
theorem finset.weighted_vsub_eq_linear_combination {k : Type u_1} {V : Type u_2} [ring k] [add_comm_group V] [module k V] {ι : Type u_3} (s : finset ι) {w : ι → k} {p : ι → V} (hw : s.sum w = 0) :
(s.weighted_vsub p) w = s.sum (λ (i : ι), w i p i)

Viewing a module as an affine space modelled on itself, a weighted_vsub is just a linear combination.

@[simp]
theorem finset.affine_combination_eq_linear_combination {k : Type u_1} {V : Type u_2} [ring k] [add_comm_group V] [module k V] {ι : Type u_4} (s : finset ι) (p : ι → V) (w : ι → k) (hw : s.sum (λ (i : ι), w i) = 1) :
(s.affine_combination p) w = s.sum (λ (i : ι), w i p i)

Viewing a module as an affine space modelled on itself, affine combinations are just linear combinations.

@[simp]
theorem finset.affine_combination_of_eq_one_of_eq_zero {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p : ι → P) {i : ι} (his : i s) (hwi : w i = 1) (hw0 : ∀ (i2 : ι), i2 si2 iw i2 = 0) :

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.

theorem finset.affine_combination_indicator_subset {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (w : ι → k) (p : ι → P) {s₁ s₂ : finset ι} (h : s₁ s₂) :

An affine combination is unaffected by changing the weights to the corresponding indicator function and adding points to the set.

theorem finset.affine_combination_map {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} {ι₂ : Type u_5} (s₂ : finset ι₂) (e : ι₂ ι) (w : ι → k) (p : ι → P) :

An affine combination, over the image of an embedding, equals an affine combination with the same points and weights over the original finset.

theorem finset.sum_smul_vsub_eq_affine_combination_vsub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₁ p₂ : ι → P) :
s.sum (λ (i : ι), w i (p₁ i -ᵥ p₂ i)) = (s.affine_combination p₁) w -ᵥ (s.affine_combination p₂) w

A weighted sum of pairwise subtractions, expressed as a subtraction of two affine_combination expressions.

theorem finset.sum_smul_vsub_const_eq_affine_combination_vsub {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₁ : ι → P) (p₂ : P) (h : s.sum (λ (i : ι), w i) = 1) :
s.sum (λ (i : ι), w i (p₁ i -ᵥ p₂)) = (s.affine_combination p₁) w -ᵥ p₂

A weighted sum of pairwise subtractions, where the point on the right is constant and the sum of the weights is 1.

theorem finset.sum_smul_const_vsub_eq_vsub_affine_combination {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) (p₂ : ι → P) (p₁ : P) (h : s.sum (λ (i : ι), w i) = 1) :
s.sum (λ (i : ι), w i (p₁ -ᵥ p₂ i)) = p₁ -ᵥ (s.affine_combination p₂) w

A weighted sum of pairwise subtractions, where the point on the left is constant and the sum of the weights is 1.

theorem finset.eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} {v : V} {x : k} {s : set ι} {p : ι → P} {b : P} :
(∃ (fs : finset ι) (hfs : fs s) (w : ι → k) (hw : fs.sum (λ (i : ι), w i) = x), v = (fs.weighted_vsub_of_point p b) w) ∃ (fs : finset s) (w : s → k) (hw : fs.sum (λ (i : s), w i) = x), v = (fs.weighted_vsub_of_point (λ (i : s), p i) b) w

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.

theorem finset.eq_weighted_vsub_subset_iff_eq_weighted_vsub_subtype (k : Type u_1) {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} {v : V} {s : set ι} {p : ι → P} :
(∃ (fs : finset ι) (hfs : fs s) (w : ι → k) (hw : fs.sum (λ (i : ι), w i) = 0), v = (fs.weighted_vsub p) w) ∃ (fs : finset s) (w : s → k) (hw : fs.sum (λ (i : s), w i) = 0), v = (fs.weighted_vsub (λ (i : s), p i)) w

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.

theorem finset.eq_affine_combination_subset_iff_eq_affine_combination_subtype (k : Type u_1) (V : Type u_2) {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} {p0 : P} {s : set ι} {p : ι → P} :
(∃ (fs : finset ι) (hfs : fs s) (w : ι → k) (hw : fs.sum (λ (i : ι), w i) = 1), p0 = (fs.affine_combination p) w) ∃ (fs : finset s) (w : s → k) (hw : fs.sum (λ (i : s), w i) = 1), p0 = (fs.affine_combination (λ (i : s), p i)) w

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.

theorem finset.map_affine_combination {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [S : add_torsor V P] {ι : Type u_4} (s : finset ι) {V₂ : Type u_5} {P₂ : Type u_6} [add_comm_group V₂] [module k V₂] [add_torsor V₂ P₂] (p : ι → P) (w : ι → k) (hw : s.sum w = 1) (f : P →ᵃ[k] P₂) :

Affine maps commute with affine combinations.

def finset.centroid_weights (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) :
ι → k

The weights for the centroid of some points.

Equations
@[simp]
theorem finset.centroid_weights_apply (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) (i : ι) :

centroid_weights at any point.

theorem finset.centroid_weights_eq_const (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) :

centroid_weights equals a constant function.

theorem finset.sum_centroid_weights_eq_one_of_cast_card_ne_zero {k : Type u_1} [division_ring k] {ι : Type u_4} (s : finset ι) (h : (s.card) 0) :
s.sum (λ (i : ι), finset.centroid_weights k s i) = 1

The weights in the centroid sum to 1, if the number of points, converted to k, is not zero.

theorem finset.sum_centroid_weights_eq_one_of_card_ne_zero (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [char_zero k] (h : s.card 0) :
s.sum (λ (i : ι), finset.centroid_weights k s i) = 1

In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is not zero.

theorem finset.sum_centroid_weights_eq_one_of_nonempty (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [char_zero k] (h : s.nonempty) :
s.sum (λ (i : ι), finset.centroid_weights k s i) = 1

In the characteristic zero case, the weights in the centroid sum to 1 if the set is nonempty.

theorem finset.sum_centroid_weights_eq_one_of_card_eq_add_one (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [char_zero k] {n : } (h : s.card = n + 1) :
s.sum (λ (i : ι), finset.centroid_weights k s i) = 1

In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is n + 1.

noncomputable def finset.centroid (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) :
P

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
theorem finset.centroid_def (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (s : finset ι) (p : ι → P) :

The definition of the centroid.

theorem finset.centroid_univ (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] (s : finset P) :
@[simp]
theorem finset.centroid_singleton (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (p : ι → P) (i : ι) :
finset.centroid k {i} p = p i

The centroid of a single point.

theorem finset.centroid_pair (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [invertible 2] (p : ι → P) (i₁ i₂ : ι) :
finset.centroid k {i₁, i₂} p = 2⁻¹ (p i₂ -ᵥ p i₁) +ᵥ p i₁

The centroid of two points, expressed directly as adding a vector to a point.

theorem finset.centroid_pair_fin (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] [invertible 2] (p : fin 2 → P) :

The centroid of two points indexed by fin 2, expressed directly as adding a vector to the first point.

theorem finset.centroid_map (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} {ι₂ : Type u_5} (s₂ : finset ι₂) (e : ι₂ ι) (p : ι → P) :

A centroid, over the image of an embedding, equals a centroid with the same points and weights over the original finset.

noncomputable def finset.centroid_weights_indicator (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) :
ι → k

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.

theorem finset.sum_centroid_weights_indicator (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [fintype ι] :

The sum of the weights for the centroid indexed by a fintype.

theorem finset.sum_centroid_weights_indicator_eq_one_of_card_ne_zero (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [char_zero k] [fintype ι] (h : s.card 0) :

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.

theorem finset.sum_centroid_weights_indicator_eq_one_of_nonempty (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [char_zero k] [fintype ι] (h : s.nonempty) :

In the characteristic zero case, the weights in the centroid indexed by a fintype sum to 1 if the set is nonempty.

theorem finset.sum_centroid_weights_indicator_eq_one_of_card_eq_add_one (k : Type u_1) [division_ring k] {ι : Type u_4} (s : finset ι) [char_zero k] [fintype ι] {n : } (h : s.card = n + 1) :

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.

theorem finset.centroid_eq_affine_combination_fintype (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (s : finset ι) [fintype ι] (p : ι → P) :

The centroid as an affine combination over a fintype.

theorem finset.centroid_eq_centroid_image_of_inj_on (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (s : finset ι) {p : ι → P} (hi : ∀ (i : ι), i s∀ (j : ι), j sp i = p ji = j) {ps : set P} [fintype ps] (hps : ps = p '' s) :

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.

theorem finset.centroid_eq_of_inj_on_of_image_eq (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (s : finset ι) {ι₂ : Type u_5} (s₂ : finset ι₂) {p : ι → P} (hi : ∀ (i : ι), i s∀ (j : ι), j sp i = p ji = j) {p₂ : ι₂ → P} (hi₂ : ∀ (i : ι₂), i s₂∀ (j : ι₂), j s₂p₂ i = p₂ ji = j) (he : p '' s = p₂ '' s₂) :

Two indexed families of points that are injective on the given finsets and with the same points in the image of those finsets have the same centroid.

theorem weighted_vsub_mem_vector_span {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} {s : finset ι} {w : ι → k} (h : s.sum (λ (i : ι), w i) = 0) (p : ι → P) :

A weighted_vsub with sum of weights 0 is in the vector_span of an indexed family.

theorem affine_combination_mem_affine_span {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [nontrivial k] {s : finset ι} {w : ι → k} (h : s.sum (λ (i : ι), w i) = 1) (p : ι → P) :

An affine_combination with sum of weights 1 is in the affine_span of an indexed family, if the underlying ring is nontrivial.

theorem mem_vector_span_iff_eq_weighted_vsub (k : Type u_1) {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} {v : V} {p : ι → P} :
v vector_span k (set.range p) ∃ (s : finset ι) (w : ι → k) (h : s.sum (λ (i : ι), w i) = 0), v = (s.weighted_vsub p) w

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.

theorem eq_affine_combination_of_mem_affine_span {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} {p1 : P} {p : ι → P} (h : p1 affine_span k (set.range p)) :
∃ (s : finset ι) (w : ι → k) (hw : s.sum (λ (i : ι), w i) = 1), p1 = (s.affine_combination p) w

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.

theorem eq_affine_combination_of_mem_affine_span_of_fintype {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [fintype ι] {p1 : P} {p : ι → P} (h : p1 affine_span k (set.range p)) :
∃ (w : ι → k) (hw : finset.univ.sum (λ (i : ι), w i) = 1), p1 = (finset.univ.affine_combination p) w
theorem mem_affine_span_iff_eq_affine_combination (k : Type u_1) (V : Type u_2) {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [nontrivial k] {p1 : P} {p : ι → P} :
p1 affine_span k (set.range p) ∃ (s : finset ι) (w : ι → k) (hw : s.sum (λ (i : ι), w i) = 1), p1 = (s.affine_combination p) w

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.

theorem mem_affine_span_iff_eq_weighted_vsub_of_point_vadd (k : Type u_1) (V : Type u_2) {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [nontrivial k] (p : ι → P) (j : ι) (q : P) :
q affine_span k (set.range p) ∃ (s : finset ι) (w : ι → k), q = (s.weighted_vsub_of_point p (p j)) w +ᵥ p j

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.

theorem affine_span_eq_affine_span_line_map_units {k : Type u_1} {V : Type u_2} {P : Type u_3} [ring k] [add_comm_group V] [module k V] [add_torsor V P] [nontrivial k] {s : set P} {p : P} (hp : p s) (w : s → kˣ) :

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.

theorem centroid_mem_affine_span_of_cast_card_ne_zero {k : Type u_1} {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} {s : finset ι} (p : ι → P) (h : (s.card) 0) :

The centroid lies in the affine span if the number of points, converted to k, is not zero.

theorem centroid_mem_affine_span_of_card_ne_zero (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [char_zero k] {s : finset ι} (p : ι → P) (h : s.card 0) :

In the characteristic zero case, the centroid lies in the affine span if the number of points is not zero.

theorem centroid_mem_affine_span_of_nonempty (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [char_zero k] {s : finset ι} (p : ι → P) (h : s.nonempty) :

In the characteristic zero case, the centroid lies in the affine span if the set is nonempty.

theorem centroid_mem_affine_span_of_card_eq_add_one (k : Type u_1) {V : Type u_2} {P : Type u_3} [division_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} [char_zero k] {s : finset ι} (p : ι → P) {n : } (h : s.card = n + 1) :

In the characteristic zero case, the centroid lies in the affine span if the number of points is n + 1.

def affine_map.weighted_vsub_of_point {k : Type u_1} {V : Type u_2} (P : Type u_3) [comm_ring k] [add_comm_group V] [module k V] [add_torsor V P] {ι : Type u_4} (s : finset ι) (w : ι → k) :
(ι → P) × P →ᵃ[k] V

A weighted sum, as an affine map on the points involved.

Equations