@@ -9,19 +9,26 @@ import algebra.invertible
9
9
/-!
10
10
# Midpoint of a segment
11
11
12
- Given two points `x` and `y` in a vector space over a ring `R` with invertible `2`,
13
- we define `midpoint R x y` to be `(⅟2:R) • (x + y)`, where `(⅟2:R)` is the inverse of `(2:R)`
14
- provided by `[invertible (2:R)]`.
12
+ ## Main definitions
15
13
16
- We prove that `z` is the midpoint of `[x, y]` if and only if `x + y = z + z`, hence `midpoint R x y`
17
- does not depend on `R`. We also prove that `midpoint x y` is linear both in `x` and `y`.
14
+ * `midpoint R x y`: midpoint of the segment `[x, y]`. We define it for `x` and `y`
15
+ in a module over a ring `R` with invertible `2`.
16
+ * `add_monoid_hom.of_map_midpoint`: construct an `add_monoid_hom` given a map `f` such that
17
+ `f` sends zero to zero and midpoints to midpoints.
18
18
19
- We do not mark lemmas other than `midpoint_self` as `@[simp]` because it is hard to tell which side
20
- is simpler.
19
+ ## Main theorems
20
+
21
+ * `midpoint_eq_iff`: `z` is the midpoint of `[x, y]` if and only if `x + y = z + z`,
22
+ * `midpoint_unique`: `midpoint R x y` does not depend on `R`;
23
+ * `midpoint x y` is linear both in `x` and `y`;
24
+ * `reflection_midpoint_left`, `reflection_midpoint_right`: `equiv.reflection (midpoint R x y)`
25
+ swaps `x` and `y`.
26
+
27
+ We do not mark most lemmas as `@[simp]` because it is hard to tell which side is simpler.
21
28
22
29
## Tags
23
30
24
- midpoint
31
+ midpoint, add_monoid_hom
25
32
-/
26
33
27
34
variables (R : Type *) {E : Type *}
@@ -36,23 +43,29 @@ def midpoint (x y : E) : E := (⅟2:R) • (x + y)
36
43
lemma midpoint_eq_iff {x y z : E} : midpoint R x y = z ↔ x + y = z + z :=
37
44
⟨λ h, h ▸ calc x + y = (2 * ⅟2 :R) • (x + y) : by rw [mul_inv_of_self, one_smul]
38
45
... = midpoint R x y + midpoint R x y : by rw [two_mul, add_smul, midpoint],
39
- λ h, by rw [midpoint, h, ← one_smul R z, ← add_smul, smul_smul, ← bit0, inv_of_mul_self]⟩
40
-
41
- variable {R}
46
+ λ h, by rw [midpoint, h, ← two_smul R z, smul_smul, inv_of_mul_self, one_smul]⟩
42
47
43
- lemma midpoint_def (x y : E) : midpoint R x y = (⅟2 :R) • (x + y) := rfl
48
+ @[simp] lemma midpoint_add_self (x y : E) : midpoint R x y + midpoint R x y = x + y :=
49
+ ((midpoint_eq_iff R).1 rfl).symm
44
50
45
51
/-- `midpoint` does not depend on the ring `R`. -/
46
52
lemma midpoint_unique (R' : Type *) [semiring R'] [invertible (2 :R')] [semimodule R' E] (x y : E) :
47
53
midpoint R x y = midpoint R' x y :=
48
54
(midpoint_eq_iff R).2 $ (midpoint_eq_iff R').1 rfl
49
55
50
56
@[simp] lemma midpoint_self (x : E) : midpoint R x x = x :=
51
- by rw [midpoint_def, smul_add, ← add_smul, ← mul_two, inv_of_mul_self, one_smul]
57
+ by rw [midpoint, smul_add, ← two_smul R, smul_smul, mul_inv_of_self, one_smul]
58
+
59
+ variable {R}
60
+
61
+ lemma midpoint_def (x y : E) : midpoint R x y = (⅟2 :R) • (x + y) := rfl
52
62
53
63
lemma midpoint_comm (x y : E) : midpoint R x y = midpoint R y x :=
54
64
by simp only [midpoint_def, add_comm]
55
65
66
+ lemma midpoint_zero_add (x y : E) : midpoint R 0 (x + y) = midpoint R x y :=
67
+ (midpoint_eq_iff R).2 $ (zero_add (x + y)).symm ▸ (midpoint_eq_iff R).1 rfl
68
+
56
69
lemma midpoint_add_add (x y x' y' : E) :
57
70
midpoint R (x + x') (y + y') = midpoint R x y + midpoint R x' y' :=
58
71
by { simp only [midpoint_def, ← smul_add, add_assoc, add_left_comm x'] }
@@ -68,6 +81,8 @@ lemma midpoint_smul_smul (c : R) (x y : E) : midpoint R (c • x) (c • y) = c
68
81
69
82
end monoid
70
83
84
+ section group
85
+
71
86
variables [ring R] [invertible (2 :R)] [add_comm_group E] [module R E]
72
87
73
88
lemma midpoint_neg_neg (x y : E) : midpoint R (-x) (-y) = -midpoint R x y :=
@@ -82,3 +97,42 @@ by rw [midpoint_sub_sub, midpoint_self]
82
97
83
98
lemma midpoint_sub_left (x y z : E) : midpoint R (x - y) (x - z) = x - midpoint R y z :=
84
99
by rw [midpoint_sub_sub, midpoint_self]
100
+
101
+ end group
102
+
103
+ namespace add_monoid_hom
104
+
105
+ variables (R) (R' : Type *) {F : Type *}
106
+ [semiring R] [invertible (2 :R)] [add_comm_monoid E] [semimodule R E]
107
+ [semiring R'] [invertible (2 :R')] [add_comm_monoid F] [semimodule R' F]
108
+
109
+ /-- A map `f : E → F` sending zero to zero and midpoints to midpoints is an `add_monoid_hom`. -/
110
+ def of_map_midpoint (f : E → F) (h0 : f 0 = 0 )
111
+ (hm : ∀ x y, f (midpoint R x y) = midpoint R' (f x) (f y)) :
112
+ E →+ F :=
113
+ { to_fun := f,
114
+ map_zero' := h0,
115
+ map_add' := λ x y, -- by rw [← midpoint_self R (x + y), ← midpoint_zero_add, hm, h0]
116
+ calc f (x + y) = f 0 + f (x + y) : by rw [h0, zero_add]
117
+ ... = midpoint R' (f 0 ) (f (x + y)) + midpoint R' (f 0 ) (f (x + y)) :
118
+ (midpoint_add_self _ _ _).symm
119
+ ... = f (midpoint R x y) + f (midpoint R x y) : by rw [← hm, midpoint_zero_add]
120
+ ... = f x + f y : by rw [hm, midpoint_add_self] }
121
+
122
+ @[simp] lemma coe_of_map_midpoint (f : E → F) (h0 : f 0 = 0 )
123
+ (hm : ∀ x y, f (midpoint R x y) = midpoint R' (f x) (f y)) :
124
+ ⇑(of_map_midpoint R R' f h0 hm) = f := rfl
125
+
126
+ end add_monoid_hom
127
+
128
+ namespace equiv
129
+
130
+ variables [ring R] [invertible (2 :R)] [add_comm_group E] [module R E]
131
+
132
+ @[simp] lemma reflection_midpoint_left (x y : E) : (reflection (midpoint R x y) : E → E) x = y :=
133
+ by rw [reflection_apply, midpoint_add_self, add_sub_cancel']
134
+
135
+ @[simp] lemma reflection_midpoint_right (x y : E) : (reflection (midpoint R x y) : E → E) y = x :=
136
+ by rw [reflection_apply, midpoint_add_self, add_sub_cancel]
137
+
138
+ end equiv
0 commit comments