Betweenness in affine spaces #
This file defines notions of a point in an affine space being between two given points.
Main definitions #
- affineSegment R x y: The segment of points weakly between- xand- y.
- Wbtw R x y z: The point- yis weakly between- xand- z.
- Sbtw R x y z: The point- yis strictly between- xand- z.
def
affineSegment
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
Set P
The segment of points weakly between x and y. When convexity is refactored to support
abstract affine combination spaces, this will no longer need to be a separate definition from
segment. However, lemmas involving +ᵥ or -ᵥ will still be relevant after such a
refactoring, as distinct from versions involving + or - in a module.
Equations
- affineSegment R x y = ⇑(AffineMap.lineMap x y) '' Set.Icc 0 1
Instances For
theorem
affineSegment_eq_segment
(R : Type u_1)
{V : Type u_2}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
(x : V)
(y : V)
 :
affineSegment R x y = segment R x y
theorem
affineSegment_comm
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
affineSegment R x y = affineSegment R y x
theorem
left_mem_affineSegment
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
x ∈ affineSegment R x y
theorem
right_mem_affineSegment
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
y ∈ affineSegment R x y
@[simp]
theorem
affineSegment_same
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
 :
affineSegment R x x = {x}
@[simp]
theorem
affineSegment_image
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
(f : P →ᵃ[R] P')
(x : P)
(y : P)
 :
⇑f '' affineSegment R x y = affineSegment R (f x) (f y)
@[simp]
theorem
affineSegment_const_vadd_image
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
(v : V)
 :
(fun (x : P) => v +ᵥ x) '' affineSegment R x y = affineSegment R (v +ᵥ x) (v +ᵥ y)
@[simp]
theorem
affineSegment_vadd_const_image
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : V)
(y : V)
(p : P)
 :
(fun (x : V) => x +ᵥ p) '' affineSegment R x y = affineSegment R (x +ᵥ p) (y +ᵥ p)
@[simp]
theorem
affineSegment_const_vsub_image
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
(p : P)
 :
(fun (x : P) => p -ᵥ x) '' affineSegment R x y = affineSegment R (p -ᵥ x) (p -ᵥ y)
@[simp]
theorem
affineSegment_vsub_const_image
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
(p : P)
 :
(fun (x : P) => x -ᵥ p) '' affineSegment R x y = affineSegment R (x -ᵥ p) (y -ᵥ p)
@[simp]
theorem
mem_const_vadd_affineSegment
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(v : V)
 :
v +ᵥ z ∈ affineSegment R (v +ᵥ x) (v +ᵥ y) ↔ z ∈ affineSegment R x y
@[simp]
theorem
mem_vadd_const_affineSegment
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : V}
{y : V}
{z : V}
(p : P)
 :
z +ᵥ p ∈ affineSegment R (x +ᵥ p) (y +ᵥ p) ↔ z ∈ affineSegment R x y
@[simp]
theorem
mem_const_vsub_affineSegment
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(p : P)
 :
p -ᵥ z ∈ affineSegment R (p -ᵥ x) (p -ᵥ y) ↔ z ∈ affineSegment R x y
@[simp]
theorem
mem_vsub_const_affineSegment
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(p : P)
 :
z -ᵥ p ∈ affineSegment R (x -ᵥ p) (y -ᵥ p) ↔ z ∈ affineSegment R x y
def
Wbtw
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
(z : P)
 :
The point y is weakly between x and z.
Equations
- Wbtw R x y z = (y ∈ affineSegment R x z)
Instances For
theorem
mem_segment_iff_wbtw
{R : Type u_1}
{V : Type u_2}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
{x : V}
{y : V}
{z : V}
 :
theorem
Function.Injective.wbtw_map_iff
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
{x : P}
{y : P}
{z : P}
{f : P →ᵃ[R] P'}
(hf : Function.Injective ⇑f)
 :
theorem
Function.Injective.sbtw_map_iff
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
{x : P}
{y : P}
{z : P}
{f : P →ᵃ[R] P'}
(hf : Function.Injective ⇑f)
 :
@[simp]
theorem
AffineEquiv.wbtw_map_iff
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
{x : P}
{y : P}
{z : P}
(f : P ≃ᵃ[R] P')
 :
@[simp]
theorem
AffineEquiv.sbtw_map_iff
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
{x : P}
{y : P}
{z : P}
(f : P ≃ᵃ[R] P')
 :
@[simp]
theorem
wbtw_const_vadd_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(v : V)
 :
@[simp]
theorem
wbtw_vadd_const_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : V}
{y : V}
{z : V}
(p : P)
 :
@[simp]
theorem
wbtw_const_vsub_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(p : P)
 :
@[simp]
theorem
wbtw_vsub_const_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(p : P)
 :
@[simp]
theorem
sbtw_const_vadd_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(v : V)
 :
@[simp]
theorem
sbtw_vadd_const_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : V}
{y : V}
{z : V}
(p : P)
 :
@[simp]
theorem
sbtw_const_vsub_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(p : P)
 :
@[simp]
theorem
sbtw_vsub_const_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(p : P)
 :
theorem
Sbtw.ne_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
y ≠ x
theorem
Sbtw.left_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
x ≠ y
theorem
Sbtw.ne_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
y ≠ z
theorem
Sbtw.right_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
z ≠ y
theorem
Sbtw.mem_image_Ioo
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
y ∈ ⇑(AffineMap.lineMap x z) '' Set.Ioo 0 1
theorem
Wbtw.mem_affineSpan
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
y ∈ affineSpan R {x, z}
@[simp]
theorem
wbtw_self_left
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
Wbtw R x x y
@[simp]
theorem
wbtw_self_right
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
Wbtw R x y y
@[simp]
theorem
wbtw_self_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
 :
@[simp]
theorem
not_sbtw_self_left
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
@[simp]
theorem
not_sbtw_self_right
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
theorem
Wbtw.left_ne_right_of_ne_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
(hne : y ≠ x)
 :
x ≠ z
theorem
Wbtw.left_ne_right_of_ne_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
(hne : y ≠ z)
 :
x ≠ z
theorem
Sbtw.left_ne_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
x ≠ z
theorem
sbtw_iff_mem_image_Ioo_and_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
 :
@[simp]
theorem
not_sbtw_self
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
theorem
wbtw_swap_left_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
(z : P)
 :
theorem
wbtw_swap_right_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
(x : P)
{y : P}
{z : P}
 :
theorem
wbtw_rotate_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
(x : P)
{y : P}
{z : P}
 :
theorem
Wbtw.swap_left_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
theorem
Wbtw.swap_right_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
theorem
Wbtw.rotate_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
theorem
Sbtw.not_swap_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
theorem
Sbtw.not_swap_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
theorem
Sbtw.not_rotate
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
@[simp]
theorem
wbtw_lineMap_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{r : R}
 :
@[simp]
theorem
sbtw_lineMap_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{x : P}
{y : P}
{r : R}
 :
@[simp]
theorem
wbtw_mul_sub_add_iff
{R : Type u_1}
[OrderedRing R]
[NoZeroDivisors R]
{x : R}
{y : R}
{r : R}
 :
@[simp]
theorem
sbtw_mul_sub_add_iff
{R : Type u_1}
[OrderedRing R]
[NoZeroDivisors R]
{x : R}
{y : R}
{r : R}
 :
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Wbtw.trans_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w y z)
(h₂ : Wbtw R w x y)
 :
Wbtw R w x z
theorem
Wbtw.trans_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w x z)
(h₂ : Wbtw R x y z)
 :
Wbtw R w y z
theorem
Wbtw.trans_sbtw_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w y z)
(h₂ : Sbtw R w x y)
 :
Sbtw R w x z
theorem
Wbtw.trans_sbtw_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w x z)
(h₂ : Sbtw R x y z)
 :
Sbtw R w y z
theorem
Sbtw.trans_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Sbtw R w y z)
(h₂ : Sbtw R w x y)
 :
Sbtw R w x z
theorem
Sbtw.trans_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Sbtw R w x z)
(h₂ : Sbtw R x y z)
 :
Sbtw R w y z
theorem
Wbtw.trans_left_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w y z)
(h₂ : Wbtw R w x y)
(h : y ≠ z)
 :
x ≠ z
theorem
Wbtw.trans_right_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w x z)
(h₂ : Wbtw R x y z)
(h : w ≠ x)
 :
w ≠ y
theorem
Sbtw.trans_wbtw_left_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Sbtw R w y z)
(h₂ : Wbtw R w x y)
 :
x ≠ z
theorem
Sbtw.trans_wbtw_right_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Sbtw R w x z)
(h₂ : Wbtw R x y z)
 :
w ≠ y
theorem
Sbtw.affineCombination_of_mem_affineSpan_pair
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[OrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroDivisors R]
[NoZeroSMulDivisors R V]
{ι : Type u_6}
{p : ι → P}
(ha : AffineIndependent R p)
{w : ι → R}
{w₁ : ι → R}
{w₂ : ι → R}
{s : Finset ι}
(hw : ∑ i ∈ s, w i = 1)
(hw₁ : ∑ i ∈ s, w₁ i = 1)
(hw₂ : ∑ i ∈ s, w₂ i = 1)
(h : (Finset.affineCombination R s p) w ∈   affineSpan R {(Finset.affineCombination R s p) w₁, (Finset.affineCombination R s p) w₂})
{i : ι}
(his : i ∈ s)
(hs : Sbtw R (w₁ i) (w i) (w₂ i))
 :
Sbtw R ((Finset.affineCombination R s p) w₁) ((Finset.affineCombination R s p) w) ((Finset.affineCombination R s p) w₂)
theorem
Wbtw.sameRay_vsub
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[StrictOrderedCommRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
theorem
Wbtw.sameRay_vsub_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[StrictOrderedCommRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
theorem
Wbtw.sameRay_vsub_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[StrictOrderedCommRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
theorem
sbtw_of_sbtw_of_sbtw_of_mem_affineSpan_pair
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedRing R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[NoZeroSMulDivisors R V]
{t : Affine.Triangle R P}
{i₁ : Fin 3}
{i₂ : Fin 3}
{i₃ : Fin 3}
(h₁₂ : i₁ ≠ i₂)
{p₁ : P}
{p₂ : P}
{p : P}
(h₁ : Sbtw R (t.points i₂) p₁ (t.points i₃))
(h₂ : Sbtw R (t.points i₁) p₂ (t.points i₃))
(h₁' : p ∈ affineSpan R {t.points i₁, p₁})
(h₂' : p ∈ affineSpan R {t.points i₂, p₂})
 :
Sbtw R (t.points i₁) p p₁
Suppose lines from two vertices of a triangle to interior points of the opposite side meet at
p. Then p lies in the interior of the first (and by symmetry the other) segment from a
vertex to the point on the opposite side.
theorem
wbtw_iff_left_eq_or_right_mem_image_Ici
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
 :
theorem
Wbtw.right_mem_image_Ici_of_left_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
(hne : x ≠ y)
 :
z ∈ ⇑(AffineMap.lineMap x y) '' Set.Ici 1
theorem
Wbtw.right_mem_affineSpan_of_left_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
(hne : x ≠ y)
 :
z ∈ affineSpan R {x, y}
theorem
sbtw_iff_left_ne_and_right_mem_image_Ioi
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
 :
theorem
Sbtw.right_mem_image_Ioi
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
z ∈ ⇑(AffineMap.lineMap x y) '' Set.Ioi 1
theorem
Sbtw.right_mem_affineSpan
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
z ∈ affineSpan R {x, y}
theorem
wbtw_iff_right_eq_or_left_mem_image_Ici
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
 :
theorem
Wbtw.left_mem_image_Ici_of_right_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
(hne : z ≠ y)
 :
x ∈ ⇑(AffineMap.lineMap z y) '' Set.Ici 1
theorem
Wbtw.left_mem_affineSpan_of_right_ne
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
(hne : z ≠ y)
 :
x ∈ affineSpan R {z, y}
theorem
sbtw_iff_right_ne_and_left_mem_image_Ioi
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
 :
theorem
Sbtw.left_mem_image_Ioi
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
x ∈ ⇑(AffineMap.lineMap z y) '' Set.Ioi 1
theorem
Sbtw.left_mem_affineSpan
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Sbtw R x y z)
 :
x ∈ affineSpan R {z, y}
theorem
Wbtw.trans_left_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w y z)
(h₂ : Wbtw R w x y)
 :
Wbtw R x y z
theorem
Wbtw.trans_right_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Wbtw R w x z)
(h₂ : Wbtw R x y z)
 :
Wbtw R w x y
theorem
Sbtw.trans_left_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Sbtw R w y z)
(h₂ : Sbtw R w x y)
 :
Sbtw R x y z
theorem
Sbtw.trans_right_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{w : P}
{x : P}
{y : P}
{z : P}
(h₁ : Sbtw R w x z)
(h₂ : Sbtw R x y z)
 :
Sbtw R w x y
theorem
Wbtw.collinear
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Wbtw R x y z)
 :
Collinear R {x, y, z}
theorem
Collinear.wbtw_or_wbtw_or_wbtw
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
(h : Collinear R {x, y, z})
 :
theorem
wbtw_iff_sameRay_vsub
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
{z : P}
 :
theorem
wbtw_pointReflection
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
Wbtw R y x ((AffineEquiv.pointReflection R x) y)
theorem
sbtw_pointReflection_of_ne
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
(h : x ≠ y)
 :
Sbtw R y x ((AffineEquiv.pointReflection R x) y)
theorem
wbtw_midpoint
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
(y : P)
 :
theorem
sbtw_midpoint_of_ne
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[LinearOrderedField R]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x : P}
{y : P}
(h : x ≠ y)
 :