doc-next-gen

Mathlib.Analysis.Normed.Affine.AddTorsor

Module docstring

{"# Torsors of normed space actions.

This file contains lemmas about normed additive torsors over normed spaces. "}

AffineSubspace.isClosed_direction_iff theorem
(s : AffineSubspace π•œ Q) : IsClosed (s.direction : Set W) ↔ IsClosed (s : Set Q)
Full source
theorem AffineSubspace.isClosed_direction_iff (s : AffineSubspace π•œ Q) :
    IsClosedIsClosed (s.direction : Set W) ↔ IsClosed (s : Set Q) := by
  rcases s.eq_bot_or_nonempty with (rfl | ⟨x, hx⟩); · simp [isClosed_singleton]
  rw [← (IsometryEquiv.vaddConst x).toHomeomorph.symm.isClosed_image,
    AffineSubspace.coe_direction_eq_vsub_set_right hx]
  rfl
Closedness of Affine Subspace Equivalent to Closedness of Its Direction
Informal description
For any affine subspace $s$ of a normed additive torsor $Q$ over a normed space $\mathbb{K}$, the direction subspace $W$ of $s$ is closed if and only if $s$ itself is closed as a subset of $Q$. In other words, the following equivalence holds: \[ \text{IsClosed}(s.\text{direction}) \leftrightarrow \text{IsClosed}(s) \] where $s.\text{direction}$ denotes the direction vector space of the affine subspace $s$.
dist_center_homothety theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist p₁ (homothety p₁ c pβ‚‚) = β€–cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_center_homothety (p₁ pβ‚‚ : P) (c : π•œ) :
    dist p₁ (homothety p₁ c pβ‚‚) = β€–cβ€– * dist p₁ pβ‚‚ := by
  simp [homothety_def, norm_smul, ← dist_eq_norm_vsub, dist_comm]
Distance from Center to Homothety Point Equals Scaled Distance
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between $p_1$ and the homothety of $p_2$ centered at $p_1$ with scale factor $c$ is equal to the product of the norm of $c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(p_1, \text{homothety}(p_1, c, p_2)) = \|c\| \cdot \text{dist}(p_1, p_2). \]
nndist_center_homothety theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist p₁ (homothety p₁ c pβ‚‚) = β€–cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_center_homothety (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist p₁ (homothety p₁ c pβ‚‚) = β€–cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_center_homothety _ _ _
Nonnegative Distance from Center to Homothety Point Equals Scaled Nonnegative Distance
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between $p_1$ and the homothety of $p_2$ centered at $p_1$ with scale factor $c$ is equal to the product of the nonnegative norm of $c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(p_1, \text{homothety}(p_1, c, p_2)) = \|c\|_+ \cdot \text{nndist}(p_1, p_2). \]
dist_homothety_center theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist (homothety p₁ c pβ‚‚) p₁ = β€–cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_homothety_center (p₁ pβ‚‚ : P) (c : π•œ) :
    dist (homothety p₁ c pβ‚‚) p₁ = β€–cβ€– * dist p₁ pβ‚‚ := by rw [dist_comm, dist_center_homothety]
Distance from Homothety Point to Center Equals Scaled Distance
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between the homothety of $p_2$ centered at $p_1$ with scale factor $c$ and $p_1$ is equal to the product of the norm of $c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(\text{homothety}(p_1, c, p_2), p_1) = \|c\| \cdot \text{dist}(p_1, p_2). \]
nndist_homothety_center theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist (homothety p₁ c pβ‚‚) p₁ = β€–cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_homothety_center (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist (homothety p₁ c pβ‚‚) p₁ = β€–cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_homothety_center _ _ _
Nonnegative Distance from Homothety Point to Center Equals Scaled Nonnegative Distance
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between the homothety of $p_2$ centered at $p_1$ with scale factor $c$ and $p_1$ is equal to the product of the nonnegative norm of $c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(\text{homothety}(p_1, c, p_2), p_1) = \|c\|_+ \cdot \text{nndist}(p_1, p_2). \]
dist_lineMap_lineMap theorem
(p₁ pβ‚‚ : P) (c₁ cβ‚‚ : π•œ) : dist (lineMap p₁ pβ‚‚ c₁) (lineMap p₁ pβ‚‚ cβ‚‚) = dist c₁ cβ‚‚ * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_lineMap_lineMap (p₁ pβ‚‚ : P) (c₁ cβ‚‚ : π•œ) :
    dist (lineMap p₁ pβ‚‚ c₁) (lineMap p₁ pβ‚‚ cβ‚‚) = dist c₁ cβ‚‚ * dist p₁ pβ‚‚ := by
  rw [dist_comm p₁ pβ‚‚]
  simp only [lineMap_apply, dist_eq_norm_vsub, vadd_vsub_vadd_cancel_right,
    ← sub_smul, norm_smul, vsub_eq_sub]
Distance Between Affine Combinations in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalars $c_1, c_2 \in \mathbb{K}$, the distance between the affine combinations $\text{lineMap}(p_1, p_2, c_1)$ and $\text{lineMap}(p_1, p_2, c_2)$ is equal to the product of the distance between $c_1$ and $c_2$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(\text{lineMap}(p_1, p_2, c_1), \text{lineMap}(p_1, p_2, c_2)) = \text{dist}(c_1, c_2) \cdot \text{dist}(p_1, p_2). \]
nndist_lineMap_lineMap theorem
(p₁ pβ‚‚ : P) (c₁ cβ‚‚ : π•œ) : nndist (lineMap p₁ pβ‚‚ c₁) (lineMap p₁ pβ‚‚ cβ‚‚) = nndist c₁ cβ‚‚ * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_lineMap_lineMap (p₁ pβ‚‚ : P) (c₁ cβ‚‚ : π•œ) :
    nndist (lineMap p₁ pβ‚‚ c₁) (lineMap p₁ pβ‚‚ cβ‚‚) = nndist c₁ cβ‚‚ * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_lineMap_lineMap _ _ _ _
Nonnegative Distance Between Affine Combinations in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalars $c_1, c_2 \in \mathbb{K}$, the nonnegative distance between the affine combinations $\text{lineMap}(p_1, p_2, c_1)$ and $\text{lineMap}(p_1, p_2, c_2)$ is equal to the product of the nonnegative distance between $c_1$ and $c_2$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(\text{lineMap}(p_1, p_2, c_1), \text{lineMap}(p_1, p_2, c_2)) = \text{nndist}(c_1, c_2) \cdot \text{nndist}(p_1, p_2). \]
lipschitzWith_lineMap theorem
(p₁ pβ‚‚ : P) : LipschitzWith (nndist p₁ pβ‚‚) (lineMap p₁ pβ‚‚ : π•œ β†’ P)
Full source
theorem lipschitzWith_lineMap (p₁ pβ‚‚ : P) : LipschitzWith (nndist p₁ pβ‚‚) (lineMap p₁ pβ‚‚ : π•œ β†’ P) :=
  LipschitzWith.of_dist_le_mul fun c₁ cβ‚‚ =>
    ((dist_lineMap_lineMap p₁ pβ‚‚ c₁ cβ‚‚).trans (mul_comm _ _)).le
Lipschitz Continuity of Affine Combination Map with Constant Equal to Point Distance
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the affine map $\text{lineMap}(p_1, p_2) : \mathbb{K} \to P$ is Lipschitz continuous with Lipschitz constant equal to the nonnegative distance between $p_1$ and $p_2$. That is, the map satisfies the Lipschitz condition: \[ \text{LipschitzWith}(\text{nndist}(p_1, p_2)) (\text{lineMap}(p_1, p_2)). \]
dist_lineMap_left theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist (lineMap p₁ pβ‚‚ c) p₁ = β€–cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_lineMap_left (p₁ pβ‚‚ : P) (c : π•œ) : dist (lineMap p₁ pβ‚‚ c) p₁ = β€–cβ€– * dist p₁ pβ‚‚ := by
  simpa only [lineMap_apply_zero, dist_zero_right] using dist_lineMap_lineMap p₁ pβ‚‚ c 0
Distance from Affine Combination to First Point in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between the affine combination $\text{lineMap}(p_1, p_2, c)$ and $p_1$ is equal to the product of the norm of $c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(\text{lineMap}(p_1, p_2, c), p_1) = \|c\| \cdot \text{dist}(p_1, p_2). \]
nndist_lineMap_left theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist (lineMap p₁ pβ‚‚ c) p₁ = β€–cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_lineMap_left (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist (lineMap p₁ pβ‚‚ c) p₁ = β€–cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_lineMap_left _ _ _
Nonnegative Distance from Affine Combination to First Point in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between the affine combination $\text{lineMap}(p_1, p_2, c)$ and $p_1$ is equal to the product of the seminorm of $c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(\text{lineMap}(p_1, p_2, c), p_1) = \|c\|_+ \cdot \text{nndist}(p_1, p_2). \]
dist_left_lineMap theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist p₁ (lineMap p₁ pβ‚‚ c) = β€–cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_left_lineMap (p₁ pβ‚‚ : P) (c : π•œ) : dist p₁ (lineMap p₁ pβ‚‚ c) = β€–cβ€– * dist p₁ pβ‚‚ :=
  (dist_comm _ _).trans (dist_lineMap_left _ _ _)
Distance from Point to Affine Combination in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between $p_1$ and the affine combination $\text{lineMap}(p_1, p_2, c)$ is equal to the product of the norm of $c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(p_1, \text{lineMap}(p_1, p_2, c)) = \|c\| \cdot \text{dist}(p_1, p_2). \]
nndist_left_lineMap theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist p₁ (lineMap p₁ pβ‚‚ c) = β€–cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_left_lineMap (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist p₁ (lineMap p₁ pβ‚‚ c) = β€–cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_left_lineMap _ _ _
Nonnegative Distance from Point to Affine Combination in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between $p_1$ and the affine combination $\text{lineMap}(p_1, p_2, c)$ is equal to the product of the nonnegative norm of $c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(p_1, \text{lineMap}(p_1, p_2, c)) = \|c\|_+ \cdot \text{nndist}(p_1, p_2). \]
dist_lineMap_right theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist (lineMap p₁ pβ‚‚ c) pβ‚‚ = β€–1 - cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_lineMap_right (p₁ pβ‚‚ : P) (c : π•œ) :
    dist (lineMap p₁ pβ‚‚ c) pβ‚‚ = β€–1 - cβ€– * dist p₁ pβ‚‚ := by
  simpa only [lineMap_apply_one, dist_eq_norm'] using dist_lineMap_lineMap p₁ pβ‚‚ c 1
Distance from Affine Combination to Endpoint in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between the affine combination $\text{lineMap}(p_1, p_2, c)$ and $p_2$ is equal to the product of the norm of $1 - c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(\text{lineMap}(p_1, p_2, c), p_2) = \|1 - c\| \cdot \text{dist}(p_1, p_2). \]
nndist_lineMap_right theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist (lineMap p₁ pβ‚‚ c) pβ‚‚ = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_lineMap_right (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist (lineMap p₁ pβ‚‚ c) pβ‚‚ = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_lineMap_right _ _ _
Nonnegative Distance from Affine Combination to Endpoint in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between the affine combination $\text{lineMap}(p_1, p_2, c)$ and $p_2$ is equal to the product of the nonnegative norm of $1 - c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(\text{lineMap}(p_1, p_2, c), p_2) = \|1 - c\|_+ \cdot \text{nndist}(p_1, p_2). \]
dist_right_lineMap theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist pβ‚‚ (lineMap p₁ pβ‚‚ c) = β€–1 - cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_right_lineMap (p₁ pβ‚‚ : P) (c : π•œ) : dist pβ‚‚ (lineMap p₁ pβ‚‚ c) = β€–1 - cβ€– * dist p₁ pβ‚‚ :=
  (dist_comm _ _).trans (dist_lineMap_right _ _ _)
Distance from Endpoint to Affine Combination in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between $p_2$ and the affine combination $\text{lineMap}(p_1, p_2, c)$ is equal to the product of the norm of $1 - c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(p_2, \text{lineMap}(p_1, p_2, c)) = \|1 - c\| \cdot \text{dist}(p_1, p_2). \]
nndist_right_lineMap theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist pβ‚‚ (lineMap p₁ pβ‚‚ c) = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_right_lineMap (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist pβ‚‚ (lineMap p₁ pβ‚‚ c) = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_right_lineMap _ _ _
Nonnegative Distance from Endpoint to Affine Combination in Normed Torsor
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between $p_2$ and the affine combination $\text{lineMap}(p_1, p_2, c)$ is equal to the product of the nonnegative norm of $1 - c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(p_2, \text{lineMap}(p_1, p_2, c)) = \|1 - c\|_+ \cdot \text{nndist}(p_1, p_2). \]
dist_homothety_self theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist (homothety p₁ c pβ‚‚) pβ‚‚ = β€–1 - cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_homothety_self (p₁ pβ‚‚ : P) (c : π•œ) :
    dist (homothety p₁ c pβ‚‚) pβ‚‚ = β€–1 - cβ€– * dist p₁ pβ‚‚ := by
  rw [homothety_eq_lineMap, dist_lineMap_right]
Distance from Homothety to Endpoint in Normed Torsor: $\text{dist}(\text{homothety}(p_1, c, p_2), p_2) = \|1 - c\| \cdot \text{dist}(p_1, p_2)$
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance between the homothety $\text{homothety}(p_1, c, p_2)$ and $p_2$ is equal to the product of the norm of $1 - c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(\text{homothety}(p_1, c, p_2), p_2) = \|1 - c\| \cdot \text{dist}(p_1, p_2). \]
nndist_homothety_self theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist (homothety p₁ c pβ‚‚) pβ‚‚ = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_homothety_self (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist (homothety p₁ c pβ‚‚) pβ‚‚ = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_homothety_self _ _ _
Nonnegative Distance Identity for Homothety: $\text{nndist}(\text{homothety}(p_1, c, p_2), p_2) = \|1 - c\|_{\mathbb{R}_{\geq 0}} \cdot \text{nndist}(p_1, p_2)$
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance between the homothety $\text{homothety}(p_1, c, p_2)$ and $p_2$ is equal to the product of the nonnegative norm of $1 - c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(\text{homothety}(p_1, c, p_2), p_2) = \|1 - c\|_{\mathbb{R}_{\geq 0}} \cdot \text{nndist}(p_1, p_2). \]
dist_self_homothety theorem
(p₁ pβ‚‚ : P) (c : π•œ) : dist pβ‚‚ (homothety p₁ c pβ‚‚) = β€–1 - cβ€– * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_self_homothety (p₁ pβ‚‚ : P) (c : π•œ) :
    dist pβ‚‚ (homothety p₁ c pβ‚‚) = β€–1 - cβ€– * dist p₁ pβ‚‚ := by rw [dist_comm, dist_homothety_self]
Distance from Endpoint to Homothety in Normed Torsor: $\text{dist}(p_2, \text{homothety}(p_1, c, p_2)) = \|1 - c\| \cdot \text{dist}(p_1, p_2)$
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the distance from $p_2$ to the homothety $\text{homothety}(p_1, c, p_2)$ is equal to the product of the norm of $1 - c$ and the distance between $p_1$ and $p_2$. That is, \[ \text{dist}(p_2, \text{homothety}(p_1, c, p_2)) = \|1 - c\| \cdot \text{dist}(p_1, p_2). \]
nndist_self_homothety theorem
(p₁ pβ‚‚ : P) (c : π•œ) : nndist pβ‚‚ (homothety p₁ c pβ‚‚) = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_self_homothety (p₁ pβ‚‚ : P) (c : π•œ) :
    nndist pβ‚‚ (homothety p₁ c pβ‚‚) = β€–1 - cβ€–β‚Š * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_self_homothety _ _ _
Nonnegative Distance Identity for Homothety: $\text{nndist}(p_2, \text{homothety}(p_1, c, p_2)) = \|1 - c\|_{\mathbb{R}_{\geq 0}} \cdot \text{nndist}(p_1, p_2)$
Informal description
For any points $p_1, p_2$ in a normed additive torsor $P$ over a normed space $\mathbb{K}$, and any scalar $c \in \mathbb{K}$, the nonnegative distance from $p_2$ to the homothety $\text{homothety}(p_1, c, p_2)$ is equal to the product of the nonnegative norm of $1 - c$ and the nonnegative distance between $p_1$ and $p_2$. That is, \[ \text{nndist}(p_2, \text{homothety}(p_1, c, p_2)) = \|1 - c\|_{\mathbb{R}_{\geq 0}} \cdot \text{nndist}(p_1, p_2). \]
dist_left_midpoint theorem
(p₁ pβ‚‚ : P) : dist p₁ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_left_midpoint (p₁ pβ‚‚ : P) : dist p₁ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)β€–β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚ := by
  rw [midpoint, dist_comm, dist_lineMap_left, invOf_eq_inv, ← norm_inv]
Distance from Point to Midpoint in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the distance from $p_1$ to the midpoint of $p_1$ and $p_2$ is equal to the distance between $p_1$ and $p_2$ divided by the norm of $2$ in $\mathbb{K}$. That is, \[ \text{dist}(p_1, \text{midpoint}_{\mathbb{K}}(p_1, p_2)) = \|2\|^{-1} \cdot \text{dist}(p_1, p_2). \]
nndist_left_midpoint theorem
(p₁ pβ‚‚ : P) : nndist p₁ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_left_midpoint (p₁ pβ‚‚ : P) :
    nndist p₁ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)β€–β‚Šβ€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_left_midpoint _ _
Nonnegative Distance from Point to Midpoint in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the nonnegative distance from $p_1$ to the midpoint of $p_1$ and $p_2$ is equal to the nonnegative distance between $p_1$ and $p_2$ divided by the nonnegative norm of $2$ in $\mathbb{K}$. That is, \[ \text{nndist}(p_1, \text{midpoint}_{\mathbb{K}}(p_1, p_2)) = \|2\|_{\text{nn}}^{-1} \cdot \text{nndist}(p_1, p_2). \]
dist_midpoint_left theorem
(p₁ pβ‚‚ : P) : dist (midpoint π•œ p₁ pβ‚‚) p₁ = β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_midpoint_left (p₁ pβ‚‚ : P) : dist (midpoint π•œ p₁ pβ‚‚) p₁ = β€–(2 : π•œ)β€–β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚ := by
  rw [dist_comm, dist_left_midpoint]
Distance from Midpoint to Left Point in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the distance from the midpoint of $p_1$ and $p_2$ to $p_1$ is equal to the distance between $p_1$ and $p_2$ divided by the norm of $2$ in $\mathbb{K}$. That is, \[ \text{dist}(\text{midpoint}_{\mathbb{K}}(p_1, p_2), p_1) = \|2\|^{-1} \cdot \text{dist}(p_1, p_2). \]
nndist_midpoint_left theorem
(p₁ pβ‚‚ : P) : nndist (midpoint π•œ p₁ pβ‚‚) p₁ = β€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_midpoint_left (p₁ pβ‚‚ : P) :
    nndist (midpoint π•œ p₁ pβ‚‚) p₁ = β€–(2 : π•œ)β€–β‚Šβ€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_midpoint_left _ _
Nonnegative Distance from Midpoint to Left Point in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the nonnegative distance from the midpoint of $p_1$ and $p_2$ to $p_1$ is equal to the nonnegative distance between $p_1$ and $p_2$ divided by the nonnegative norm of $2$ in $\mathbb{K}$. That is, \[ \text{nndist}(\text{midpoint}_{\mathbb{K}}(p_1, p_2), p_1) = \|2\|_{\text{nn}}^{-1} \cdot \text{nndist}(p_1, p_2). \]
dist_midpoint_right theorem
(p₁ pβ‚‚ : P) : dist (midpoint π•œ p₁ pβ‚‚) pβ‚‚ = β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_midpoint_right (p₁ pβ‚‚ : P) :
    dist (midpoint π•œ p₁ pβ‚‚) pβ‚‚ = β€–(2 : π•œ)β€–β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚ := by
  rw [midpoint_comm, dist_midpoint_left, dist_comm]
Distance from Midpoint to Right Point in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the distance from the midpoint of $p_1$ and $p_2$ to $p_2$ is equal to the distance between $p_1$ and $p_2$ divided by the norm of $2$ in $\mathbb{K}$. That is, \[ \text{dist}(\text{midpoint}_{\mathbb{K}}(p_1, p_2), p_2) = \|2\|^{-1} \cdot \text{dist}(p_1, p_2). \]
nndist_midpoint_right theorem
(p₁ pβ‚‚ : P) : nndist (midpoint π•œ p₁ pβ‚‚) pβ‚‚ = β€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_midpoint_right (p₁ pβ‚‚ : P) :
    nndist (midpoint π•œ p₁ pβ‚‚) pβ‚‚ = β€–(2 : π•œ)β€–β‚Šβ€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_midpoint_right _ _
Nonnegative Distance from Midpoint to Right Point in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the nonnegative distance from the midpoint of $p_1$ and $p_2$ to $p_2$ is equal to the nonnegative distance between $p_1$ and $p_2$ divided by the nonnegative norm of $2$ in $\mathbb{K}$. That is, \[ \text{nndist}(\text{midpoint}_{\mathbb{K}}(p_1, p_2), p_2) = \|2\|_{\text{nn}}^{-1} \cdot \text{nndist}(p_1, p_2). \]
dist_right_midpoint theorem
(p₁ pβ‚‚ : P) : dist pβ‚‚ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚
Full source
@[simp]
theorem dist_right_midpoint (p₁ pβ‚‚ : P) :
    dist pβ‚‚ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)β€–β€–(2 : π•œ)‖⁻¹ * dist p₁ pβ‚‚ := by
  rw [dist_comm, dist_midpoint_right]
Distance from Right Point to Midpoint in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the distance from $p_2$ to the midpoint of $p_1$ and $p_2$ is equal to the distance between $p_1$ and $p_2$ divided by the norm of $2$ in $\mathbb{K}$. That is, \[ \text{dist}(p_2, \text{midpoint}_{\mathbb{K}}(p_1, p_2)) = \|2\|^{-1} \cdot \text{dist}(p_1, p_2). \]
nndist_right_midpoint theorem
(p₁ pβ‚‚ : P) : nndist pβ‚‚ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚
Full source
@[simp]
theorem nndist_right_midpoint (p₁ pβ‚‚ : P) :
    nndist pβ‚‚ (midpoint π•œ p₁ pβ‚‚) = β€–(2 : π•œ)β€–β‚Šβ€–(2 : π•œ)β€–β‚Šβ»ΒΉ * nndist p₁ pβ‚‚ :=
  NNReal.eq <| dist_right_midpoint _ _
Nonnegative Distance from Right Point to Midpoint in Normed Torsor
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the nonnegative distance from $p_2$ to the midpoint of $p_1$ and $p_2$ is equal to the nonnegative distance between $p_1$ and $p_2$ divided by the nonnegative norm of $2$ in $\mathbb{K}$. That is, \[ \text{nndist}(p_2, \text{midpoint}_{\mathbb{K}}(p_1, p_2)) = \|2\|_{\text{nn}}^{-1} \cdot \text{nndist}(p_1, p_2). \]
dist_left_midpoint_eq_dist_right_midpoint theorem
(p₁ pβ‚‚ : P) : dist p₁ (midpoint π•œ p₁ pβ‚‚) = dist pβ‚‚ (midpoint π•œ p₁ pβ‚‚)
Full source
/-- The midpoint of the segment AB is the same distance from A as it is from B. -/
theorem dist_left_midpoint_eq_dist_right_midpoint (p₁ pβ‚‚ : P) :
    dist p₁ (midpoint π•œ p₁ pβ‚‚) = dist pβ‚‚ (midpoint π•œ p₁ pβ‚‚) := by
  rw [dist_left_midpoint p₁ pβ‚‚, dist_right_midpoint p₁ pβ‚‚]
Midpoint Equidistance Property in Normed Torsors
Informal description
For any points $p_1$ and $p_2$ in a normed additive torsor $P$ over a normed field $\mathbb{K}$, the distance from $p_1$ to the midpoint of $p_1$ and $p_2$ is equal to the distance from $p_2$ to the same midpoint. That is, \[ \text{dist}(p_1, \text{midpoint}_{\mathbb{K}}(p_1, p_2)) = \text{dist}(p_2, \text{midpoint}_{\mathbb{K}}(p_1, p_2)). \]
dist_midpoint_midpoint_le' theorem
(p₁ pβ‚‚ p₃ pβ‚„ : P) : dist (midpoint π•œ p₁ pβ‚‚) (midpoint π•œ p₃ pβ‚„) ≀ (dist p₁ p₃ + dist pβ‚‚ pβ‚„) / β€–(2 : π•œ)β€–
Full source
theorem dist_midpoint_midpoint_le' (p₁ pβ‚‚ p₃ pβ‚„ : P) :
    dist (midpoint π•œ p₁ pβ‚‚) (midpoint π•œ p₃ pβ‚„) ≀ (dist p₁ p₃ + dist pβ‚‚ pβ‚„) / β€–(2 : π•œ)β€– := by
  rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, midpoint_vsub_midpoint]
  rw [midpoint_eq_smul_add, norm_smul, invOf_eq_inv, norm_inv, ← div_eq_inv_mul]
  exact div_le_div_of_nonneg_right (norm_add_le _ _) (norm_nonneg _)
Distance Between Midpoints Bounded by Average of Endpoint Distances
Informal description
For any four points $p_1, p_2, p_3, p_4$ in a normed torsor $P$ over a normed space $\mathbb{K}$, the distance between the midpoints of $p_1p_2$ and $p_3p_4$ is bounded by the average of the distances $p_1p_3$ and $p_2p_4$, scaled by the norm of $2$ in $\mathbb{K}$. That is, \[ \text{dist}(\text{midpoint}(p_1, p_2), \text{midpoint}(p_3, p_4)) \leq \frac{\text{dist}(p_1, p_3) + \text{dist}(p_2, p_4)}{\|2\|}. \]
nndist_midpoint_midpoint_le' theorem
(p₁ pβ‚‚ p₃ pβ‚„ : P) : nndist (midpoint π•œ p₁ pβ‚‚) (midpoint π•œ p₃ pβ‚„) ≀ (nndist p₁ p₃ + nndist pβ‚‚ pβ‚„) / β€–(2 : π•œ)β€–β‚Š
Full source
theorem nndist_midpoint_midpoint_le' (p₁ pβ‚‚ p₃ pβ‚„ : P) :
    nndist (midpoint π•œ p₁ pβ‚‚) (midpoint π•œ p₃ pβ‚„) ≀ (nndist p₁ p₃ + nndist pβ‚‚ pβ‚„) / β€–(2 : π•œ)β€–β‚Š :=
  dist_midpoint_midpoint_le' _ _ _ _
Nonnegative Distance Between Midpoints Bounded by Average of Endpoint Distances
Informal description
For any four points $p_1, p_2, p_3, p_4$ in a normed torsor $P$ over a normed space $\mathbb{K}$, the nonnegative distance between the midpoints of $p_1p_2$ and $p_3p_4$ is bounded by the average of the nonnegative distances $p_1p_3$ and $p_2p_4$, scaled by the seminorm of $2$ in $\mathbb{K}$. That is, \[ \text{nndist}(\text{midpoint}(p_1, p_2), \text{midpoint}(p_3, p_4)) \leq \frac{\text{nndist}(p_1, p_3) + \text{nndist}(p_2, p_4)}{\|2\|_\mathbb{K}}. \]
dist_pointReflection_left theorem
(p q : P) : dist (Equiv.pointReflection p q) p = dist p q
Full source
@[simp] theorem dist_pointReflection_left (p q : P) :
    dist (Equiv.pointReflection p q) p = dist p q := by
  simp [dist_eq_norm_vsub V, Equiv.pointReflection_vsub_left (G := V)]
Distance Between Point Reflection and Center Equals Original Distance
Informal description
For any points $p$ and $q$ in a normed torsor $P$ over a normed space $\mathbb{K}$, the distance between the point reflection of $q$ about $p$ and $p$ itself is equal to the distance between $p$ and $q$, i.e., \[ \text{dist}(\text{pointReflection}(p, q), p) = \text{dist}(p, q). \]
dist_left_pointReflection theorem
(p q : P) : dist p (Equiv.pointReflection p q) = dist p q
Full source
@[simp] theorem dist_left_pointReflection (p q : P) :
    dist p (Equiv.pointReflection p q) = dist p q :=
  (dist_comm _ _).trans (dist_pointReflection_left _ _)
Distance Between Point and Its Reflection Equals Original Distance
Informal description
For any points $p$ and $q$ in a normed torsor $P$ over a normed space $\mathbb{K}$, the distance between $p$ and the point reflection of $q$ about $p$ is equal to the distance between $p$ and $q$, i.e., \[ \text{dist}(p, \text{pointReflection}(p, q)) = \text{dist}(p, q). \]
dist_pointReflection_right theorem
(p q : P) : dist (Equiv.pointReflection p q) q = β€–(2 : π•œ)β€– * dist p q
Full source
theorem dist_pointReflection_right (p q : P) :
    dist (Equiv.pointReflection p q) q = β€–(2 : π•œ)β€– * dist p q := by
  simp [dist_eq_norm_vsub V, Equiv.pointReflection_vsub_right (G := V), ← Nat.cast_smul_eq_nsmul π•œ,
    norm_smul]
Distance Between Point Reflection and Original Point in Normed Torsor
Informal description
For any points $p$ and $q$ in a normed torsor $P$ over a normed space $\mathbb{K}$, the distance between the point reflection of $q$ about $p$ and $q$ itself is equal to the norm of $2$ in $\mathbb{K}$ multiplied by the distance between $p$ and $q$, i.e., \[ \text{dist}(\text{pointReflection}(p, q), q) = \|2\|_{\mathbb{K}} \cdot \text{dist}(p, q). \]
dist_right_pointReflection theorem
(p q : P) : dist q (Equiv.pointReflection p q) = β€–(2 : π•œ)β€– * dist p q
Full source
theorem dist_right_pointReflection (p q : P) :
    dist q (Equiv.pointReflection p q) = β€–(2 : π•œ)β€– * dist p q :=
  (dist_comm _ _).trans (dist_pointReflection_right π•œ _ _)
Distance Between Point and Its Reflection Equals Twice Original Distance
Informal description
For any points $p$ and $q$ in a normed torsor $P$ over a normed space $\mathbb{K}$, the distance between $q$ and its reflection about $p$ is equal to the norm of $2$ in $\mathbb{K}$ multiplied by the distance between $p$ and $q$, i.e., \[ \text{dist}(q, \text{pointReflection}(p, q)) = \|2\|_{\mathbb{K}} \cdot \text{dist}(p, q). \]
antilipschitzWith_lineMap theorem
{p₁ pβ‚‚ : Q} (h : p₁ β‰  pβ‚‚) : AntilipschitzWith (nndist p₁ pβ‚‚)⁻¹ (lineMap p₁ pβ‚‚ : π•œ β†’ Q)
Full source
theorem antilipschitzWith_lineMap {p₁ pβ‚‚ : Q} (h : p₁ β‰  pβ‚‚) :
    AntilipschitzWith (nndist p₁ pβ‚‚)⁻¹ (lineMap p₁ pβ‚‚ : π•œ β†’ Q) :=
  AntilipschitzWith.of_le_mul_dist fun c₁ cβ‚‚ => by
    rw [dist_lineMap_lineMap, NNReal.coe_inv, ← dist_nndist, mul_left_comm,
      inv_mul_cancelβ‚€ (dist_ne_zero.2 h), mul_one]
Antilipschitz Property of Affine Maps Between Distinct Points
Informal description
For any distinct points $p_1$ and $p_2$ in a normed torsor $Q$ over a normed field $\mathbb{K}$, the affine map $\text{lineMap}(p_1, p_2) : \mathbb{K} \to Q$ is antilipschitz with constant $(\text{nndist}(p_1, p_2))^{-1}$. That is, for all $c_1, c_2 \in \mathbb{K}$, \[ \text{dist}(c_1, c_2) \leq (\text{nndist}(p_1, p_2))^{-1} \cdot \text{dist}(\text{lineMap}(p_1, p_2, c_1), \text{lineMap}(p_1, p_2, c_2)). \]
eventually_homothety_mem_of_mem_interior theorem
(x : Q) {s : Set Q} {y : Q} (hy : y ∈ interior s) : βˆ€αΆ  Ξ΄ in 𝓝 (1 : π•œ), homothety x Ξ΄ y ∈ s
Full source
theorem eventually_homothety_mem_of_mem_interior (x : Q) {s : Set Q} {y : Q} (hy : y ∈ interior s) :
    βˆ€αΆ  Ξ΄ in 𝓝 (1 : π•œ), homothety x Ξ΄ y ∈ s := by
  rw [(NormedAddCommGroup.nhds_basis_norm_lt (1 : π•œ)).eventually_iff]
  rcases eq_or_ne y x with h | h
  Β· use 1
    simp [h.symm, interior_subset hy]
  have hxy : 0 < β€–y -α΅₯ xβ€– := by rwa [norm_pos_iff, vsub_ne_zero]
  obtain ⟨u, hu₁, huβ‚‚, huβ‚ƒβŸ© := mem_interior.mp hy
  obtain ⟨Ρ, hΞ΅, hyΡ⟩ := Metric.isOpen_iff.mp huβ‚‚ y hu₃
  refine ⟨Ρ / β€–y -α΅₯ xβ€–, div_pos hΞ΅ hxy, fun Ξ΄ (hΞ΄ : β€–Ξ΄ - 1β€– < Ξ΅ / β€–y -α΅₯ xβ€–) => hu₁ (hyΞ΅ ?_)⟩
  rw [lt_div_iffβ‚€ hxy, ← norm_smul, sub_smul, one_smul] at hΞ΄
  rwa [homothety_apply, Metric.mem_ball, dist_eq_norm_vsub W, vadd_vsub_eq_sub_vsub]
Local Membership of Homothety Image in Interior Set
Informal description
For any point $x$ in a normed torsor $Q$ over a normed field $\mathbb{K}$, and any subset $s \subseteq Q$ with $y$ in the interior of $s$, there exists a neighborhood of $1$ in $\mathbb{K}$ such that for all $\delta$ in this neighborhood, the homothety of $y$ centered at $x$ with ratio $\delta$ lies in $s$.
eventually_homothety_image_subset_of_finite_subset_interior theorem
(x : Q) {s : Set Q} {t : Set Q} (ht : t.Finite) (h : t βŠ† interior s) : βˆ€αΆ  Ξ΄ in 𝓝 (1 : π•œ), homothety x Ξ΄ '' t βŠ† s
Full source
theorem eventually_homothety_image_subset_of_finite_subset_interior (x : Q) {s : Set Q} {t : Set Q}
    (ht : t.Finite) (h : t βŠ† interior s) : βˆ€αΆ  Ξ΄ in 𝓝 (1 : π•œ), homothety x Ξ΄ '' t βŠ† s := by
  suffices βˆ€ y ∈ t, βˆ€αΆ  Ξ΄ in 𝓝 (1 : π•œ), homothety x Ξ΄ y ∈ s by
    simp_rw [Set.image_subset_iff]
    exact (Filter.eventually_all_finite ht).mpr this
  intro y hy
  exact eventually_homothety_mem_of_mem_interior π•œ x (h hy)
Finite Interior Subset Stability under Homothety Near Identity
Informal description
For any point $x$ in a normed torsor $Q$ over a normed field $\mathbb{K}$, any finite subset $t \subseteq Q$ contained in the interior of a subset $s \subseteq Q$, there exists a neighborhood of $1$ in $\mathbb{K}$ such that for all $\delta$ in this neighborhood, the image of $t$ under the homothety centered at $x$ with ratio $\delta$ is contained in $s$.
dist_midpoint_midpoint_le theorem
(p₁ pβ‚‚ p₃ pβ‚„ : V) : dist (midpoint ℝ p₁ pβ‚‚) (midpoint ℝ p₃ pβ‚„) ≀ (dist p₁ p₃ + dist pβ‚‚ pβ‚„) / 2
Full source
theorem dist_midpoint_midpoint_le (p₁ pβ‚‚ p₃ pβ‚„ : V) :
    dist (midpoint ℝ p₁ pβ‚‚) (midpoint ℝ p₃ pβ‚„) ≀ (dist p₁ p₃ + dist pβ‚‚ pβ‚„) / 2 := by
  simpa using dist_midpoint_midpoint_le' (π•œ := ℝ) p₁ pβ‚‚ p₃ pβ‚„
Midpoint Distance Inequality in Real Normed Space: $\text{dist}(M_{12}, M_{34}) \leq \frac{d_{13} + d_{24}}{2}$
Informal description
For any four points $p_1, p_2, p_3, p_4$ in a real normed space $V$, the distance between the midpoints of $p_1p_2$ and $p_3p_4$ is at most the average of the distances $p_1p_3$ and $p_2p_4$. That is, \[ \text{dist}(\text{midpoint}(p_1, p_2), \text{midpoint}(p_3, p_4)) \leq \frac{\text{dist}(p_1, p_3) + \text{dist}(p_2, p_4)}{2}. \]
nndist_midpoint_midpoint_le theorem
(p₁ pβ‚‚ p₃ pβ‚„ : V) : nndist (midpoint ℝ p₁ pβ‚‚) (midpoint ℝ p₃ pβ‚„) ≀ (nndist p₁ p₃ + nndist pβ‚‚ pβ‚„) / 2
Full source
theorem nndist_midpoint_midpoint_le (p₁ pβ‚‚ p₃ pβ‚„ : V) :
    nndist (midpoint ℝ p₁ pβ‚‚) (midpoint ℝ p₃ pβ‚„) ≀ (nndist p₁ p₃ + nndist pβ‚‚ pβ‚„) / 2 :=
  dist_midpoint_midpoint_le _ _ _ _
Nonnegative Midpoint Distance Inequality: $\text{nndist}(M_{12}, M_{34}) \leq \frac{d_{13} + d_{24}}{2}$
Informal description
For any four points $p_1, p_2, p_3, p_4$ in a real normed space $V$, the nonnegative distance between the midpoints of $p_1p_2$ and $p_3p_4$ is at most the average of the nonnegative distances $p_1p_3$ and $p_2p_4$. That is, \[ \text{nndist}(\text{midpoint}(p_1, p_2), \text{midpoint}(p_3, p_4)) \leq \frac{\text{nndist}(p_1, p_3) + \text{nndist}(p_2, p_4)}{2}. \]
AffineMap.ofMapMidpoint definition
(f : P β†’ Q) (h : βˆ€ x y, f (midpoint ℝ x y) = midpoint ℝ (f x) (f y)) (hfc : Continuous f) : P →ᡃ[ℝ] Q
Full source
/-- A continuous map between two normed affine spaces is an affine map provided that
it sends midpoints to midpoints. -/
def AffineMap.ofMapMidpoint (f : P β†’ Q) (h : βˆ€ x y, f (midpoint ℝ x y) = midpoint ℝ (f x) (f y))
    (hfc : Continuous f) : P →ᡃ[ℝ] Q :=
  let c := Classical.arbitrary P
  AffineMap.mk' f (↑((AddMonoidHom.ofMapMidpoint ℝ ℝ
    ((AffineEquiv.vaddConst ℝ (f <| c)).symm ∘ f ∘ AffineEquiv.vaddConst ℝ c) (by simp)
    fun x y => by simp [h]).toRealLinearMap <| by
        apply_rules [Continuous.vadd, Continuous.vsub, continuous_const, hfc.comp, continuous_id]))
    c fun p => by simp
Affine map from midpoint-preserving continuous map
Informal description
Given a continuous map \( f : P \to Q \) between two normed affine spaces \( P \) and \( Q \) over the real numbers, if \( f \) preserves midpoints (i.e., \( f(\text{midpoint}(x, y)) = \text{midpoint}(f(x), f(y)) \) for all \( x, y \in P \)), then \( f \) is an affine map. Here, \(\text{midpoint}(x, y)\) denotes the midpoint between \( x \) and \( y \).
DilationEquiv.smulTorsor definition
(c : P) {k : π•œ} (hk : k β‰  0) : E β‰ƒα΅ˆ P
Full source
/-- Scaling by an element `k` of the scalar ring as a `DilationEquiv` with ratio `β€–kβ€–β‚Š`, mapping
from a normed space to a normed torsor over that space sending `0` to `c`. -/
@[simps]
def DilationEquiv.smulTorsor (c : P) {k : π•œ} (hk : k β‰  0) : E β‰ƒα΅ˆ P where
  toFun := (k β€’ Β· +α΅₯ c)
  invFun := k⁻¹ β€’ (Β· -α΅₯ c)
  left_inv x := by simp [inv_smul_smulβ‚€ hk]
  right_inv p := by simp [smul_inv_smulβ‚€ hk]
  edist_eq' := βŸ¨β€–kβ€–β‚Š, nnnorm_ne_zero_iff.mpr hk, fun x y ↦ by
    rw [show edist (k β€’ x +α΅₯ c) (k β€’ y +α΅₯ c) = _ from (IsometryEquiv.vaddConst c).isometry ..]
    exact edist_smulβ‚€ ..⟩
Dilation equivalence by scaling and translation in a normed torsor
Informal description
For a given point \( c \) in a normed torsor \( P \) over a normed space \( E \) and a nonzero scalar \( k \) from the scalar field \( \mathbb{K} \), the function `smulTorsor` defines a dilation equivalence \( E \simeq_{\delta} P \) that scales vectors by \( k \) and translates by \( c \). Specifically, it maps \( x \in E \) to \( k \cdot x + c \in P \), with inverse mapping \( p \in P \) to \( k^{-1} \cdot (p - c) \in E \). The dilation ratio is \( \|k\| \), and the equivalence preserves distances up to this scaling factor.
DilationEquiv.smulTorsor_ratio theorem
{c : P} {k : π•œ} (hk : k β‰  0) {x y : E} (h : dist x y β‰  0) : ratio (smulTorsor c hk) = β€–kβ€–β‚Š
Full source
@[simp]
lemma DilationEquiv.smulTorsor_ratio {c : P} {k : π•œ} (hk : k β‰  0) {x y : E}
    (h : distdist x y β‰  0) : ratio (smulTorsor c hk) = β€–kβ€–β‚Š :=
  Eq.symm <| ratio_unique_of_dist_ne_zero h <| by simp [dist_eq_norm, ← smul_sub, norm_smul]
Dilation Ratio of Scaling and Translation Equivalence in Normed Torsor: $\text{ratio}(\text{smulTorsor}\,c\,hk) = \|k\|$
Informal description
For a given point $c$ in a normed torsor $P$ over a normed space $E$ and a nonzero scalar $k$ from the scalar field $\mathbb{K}$, the dilation ratio of the dilation equivalence $\text{smulTorsor}\,c\,hk : E \simeq_{\delta} P$ is equal to the norm of $k$, i.e., $\text{ratio}(\text{smulTorsor}\,c\,hk) = \|k\|$, for any $x, y \in E$ with $x \neq y$.
DilationEquiv.smulTorsor_preimage_ball theorem
{c : P} {k : π•œ} (hk : k β‰  0) : smulTorsor c hk ⁻¹' (Metric.ball c β€–kβ€–) = Metric.ball (0 : E) 1
Full source
@[simp]
lemma DilationEquiv.smulTorsor_preimage_ball {c : P} {k : π•œ} (hk : k β‰  0) :
    smulTorsorsmulTorsor c hk ⁻¹' (Metric.ball c β€–kβ€–) = Metric.ball (0 : E) 1 := by
  aesop (add simp norm_smul)
Preimage of Ball under Dilation Equivalence in Normed Torsor
Informal description
For a point $c$ in a normed torsor $P$ over a normed space $E$ and a nonzero scalar $k$ from the scalar field $\mathbb{K}$, the preimage of the open ball centered at $c$ with radius $\|k\|$ under the dilation equivalence $\text{smulTorsor}_c^k$ is the open ball centered at $0$ with radius $1$ in $E$. In other words, $$\left(\text{smulTorsor}_c^k\right)^{-1}\left(B(c, \|k\|)\right) = B(0, 1).$$