Midpoint of a segment #
Main definitions #
- midpoint R x y: midpoint of the segment- [x, y]. We define it for- xand- yin a module over a ring- Rwith invertible- 2.
- AddMonoidHom.ofMapMidpoint: construct an- AddMonoidHomgiven a map- fsuch that- fsends zero to zero and midpoints to midpoints.
Main theorems #
- midpoint_eq_iff:- zis the midpoint of- [x, y]if and only if- x + y = z + z,
- midpoint_unique:- midpoint R x ydoes not depend on- R;
- midpoint x yis linear both in- xand- y;
- pointReflection_midpoint_left,- pointReflection_midpoint_right:- Equiv.pointReflection (midpoint R x y)swaps- xand- y.
We do not mark most lemmas as @[simp] because it is hard to tell which side is simpler.
Tags #
midpoint, AddMonoidHom
def
midpoint
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
P
midpoint x y is the midpoint of the segment [x, y].
Equations
- midpoint R x y = (AffineMap.lineMap x y) ⅟2
Instances For
@[simp]
theorem
AffineMap.map_midpoint
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
(f : P →ᵃ[R] P')
(a b : P)
 :
@[simp]
theorem
AffineEquiv.map_midpoint
{R : Type u_1}
{V : Type u_2}
{V' : Type u_3}
{P : Type u_4}
{P' : Type u_5}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
[AddCommGroup V']
[Module R V']
[AddTorsor V' P']
(f : P ≃ᵃ[R] P')
(a b : P)
 :
theorem
AffineEquiv.pointReflection_midpoint_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
@[simp]
theorem
Equiv.pointReflection_midpoint_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
theorem
midpoint_comm
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
theorem
AffineEquiv.pointReflection_midpoint_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
@[simp]
theorem
Equiv.pointReflection_midpoint_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
theorem
midpoint_eq_iff
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x y z : P}
 :
@[simp]
theorem
midpoint_pointReflection_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
@[simp]
theorem
midpoint_pointReflection_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
theorem
AffineEquiv.midpoint_pointReflection_left
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
theorem
AffineEquiv.midpoint_pointReflection_right
{R : Type u_1}
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x y : P)
 :
@[simp]
theorem
midpoint_sub_left
{R : Type u_1}
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(v₁ v₂ : V)
 :
@[simp]
theorem
midpoint_sub_right
{R : Type u_1}
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(v₁ v₂ : V)
 :
@[simp]
theorem
left_sub_midpoint
{R : Type u_1}
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(v₁ v₂ : V)
 :
@[simp]
theorem
right_sub_midpoint
{R : Type u_1}
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(v₁ v₂ : V)
 :
@[simp]
theorem
midpoint_eq_left_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x y : P}
 :
@[simp]
theorem
left_eq_midpoint_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x y : P}
 :
@[simp]
theorem
midpoint_eq_right_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x y : P}
 :
@[simp]
theorem
right_eq_midpoint_iff
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x y : P}
 :
theorem
midpoint_eq_iff'
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
{x y z : P}
 :
theorem
midpoint_unique
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(R' : Type u_6)
[Ring R']
[Invertible 2]
[Module R' V]
(x y : P)
 :
midpoint does not depend on the ring R.
@[simp]
theorem
midpoint_self
(R : Type u_1)
{V : Type u_2}
{P : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
[AddTorsor V P]
(x : P)
 :
@[simp]
theorem
midpoint_add_self
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x y : V)
 :
theorem
midpoint_zero_add
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x y : V)
 :
theorem
midpoint_eq_smul_add
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x y : V)
 :
@[simp]
theorem
midpoint_self_neg
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x : V)
 :
@[simp]
theorem
midpoint_neg_self
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x : V)
 :
@[simp]
theorem
midpoint_sub_add
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x y : V)
 :
@[simp]
theorem
midpoint_add_sub
(R : Type u_1)
{V : Type u_2}
[Ring R]
[Invertible 2]
[AddCommGroup V]
[Module R V]
(x y : V)
 :
def
AddMonoidHom.ofMapMidpoint
(R : Type u_1)
(R' : Type u_2)
{E : Type u_3}
{F : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup E]
[Module R E]
[Ring R']
[Invertible 2]
[AddCommGroup F]
[Module R' F]
(f : E → F)
(h0 : f 0 = 0)
(hm : ∀ (x y : E), f (midpoint R x y) = midpoint R' (f x) (f y))
 :
A map f : E → F sending zero to zero and midpoints to midpoints is an AddMonoidHom.
Equations
- AddMonoidHom.ofMapMidpoint R R' f h0 hm = { toFun := f, map_zero' := h0, map_add' := ⋯ }
Instances For
@[simp]
theorem
AddMonoidHom.coe_ofMapMidpoint
(R : Type u_1)
(R' : Type u_2)
{E : Type u_3}
{F : Type u_4}
[Ring R]
[Invertible 2]
[AddCommGroup E]
[Module R E]
[Ring R']
[Invertible 2]
[AddCommGroup F]
[Module R' F]
(f : E → F)
(h0 : f 0 = 0)
(hm : ∀ (x y : E), f (midpoint R x y) = midpoint R' (f x) (f y))
 :