Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit da10afc

Browse files
committed
feat(*): define equiv.reflection and isometry.reflection (#2609)
Define reflection in a point and prove basic properties. It is defined both as an `equiv.perm` of an `add_comm_group` and as an `isometric` of a `normed_group`. Other changes: * rename `two_smul` to `two_smul'`, add `two_smul` using `semimodule` instead of `add_monoid.smul`; * minor review of `algebra.midpoint`; * arguments of `isometry.dist_eq` and `isometry.edist_eq` are now explicit; * rename `isometry.inv` to `isometry.right_inv`; now it takes `right_inverse` instead of `equiv`; * drop `isometry_inv_fun`: use `h.symm.isometry` instead; * pull a few more lemmas about `equiv` and `isometry` to the `isometric` namespace.
1 parent a6415d7 commit da10afc

File tree

8 files changed

+236
-55
lines changed

8 files changed

+236
-55
lines changed

src/algebra/group_power.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -83,7 +83,7 @@ by rw [succ_smul, smul_add_comm']
8383

8484
theorem pow_two (a : M) : a^2 = a * a :=
8585
show a*(a*1)=a*a, by rw mul_one
86-
theorem two_smul (a : A) : 2•a = a + a :=
86+
theorem two_smul' (a : A) : 2•a = a + a :=
8787
show a+(a+0)=a+a, by rw add_zero
8888

8989
theorem pow_add (a : M) (m n : ℕ) : a^(m + n) = a^m * a^n :=

src/algebra/midpoint.lean

Lines changed: 67 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -9,19 +9,26 @@ import algebra.invertible
99
/-!
1010
# Midpoint of a segment
1111
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
1513
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.
1818
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.
2128
2229
## Tags
2330
24-
midpoint
31+
midpoint, add_monoid_hom
2532
-/
2633

2734
variables (R : Type*) {E : Type*}
@@ -36,23 +43,29 @@ def midpoint (x y : E) : E := (⅟2:R) • (x + y)
3643
lemma midpoint_eq_iff {x y z : E} : midpoint R x y = z ↔ x + y = z + z :=
3744
⟨λ h, h ▸ calc x + y = (2 * ⅟2:R) • (x + y) : by rw [mul_inv_of_self, one_smul]
3845
... = 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]⟩
4247

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
4450

4551
/-- `midpoint` does not depend on the ring `R`. -/
4652
lemma midpoint_unique (R' : Type*) [semiring R'] [invertible (2:R')] [semimodule R' E] (x y : E) :
4753
midpoint R x y = midpoint R' x y :=
4854
(midpoint_eq_iff R).2 $ (midpoint_eq_iff R').1 rfl
4955

5056
@[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
5262

5363
lemma midpoint_comm (x y : E) : midpoint R x y = midpoint R y x :=
5464
by simp only [midpoint_def, add_comm]
5565

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+
5669
lemma midpoint_add_add (x y x' y' : E) :
5770
midpoint R (x + x') (y + y') = midpoint R x y + midpoint R x' y' :=
5871
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
6881

6982
end monoid
7083

84+
section group
85+
7186
variables [ring R] [invertible (2:R)] [add_comm_group E] [module R E]
7287

7388
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]
8297

8398
lemma midpoint_sub_left (x y z : E) : midpoint R (x - y) (x - z) = x - midpoint R y z :=
8499
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

src/algebra/module.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -71,6 +71,8 @@ theorem add_smul : (r + s) • x = r • x + s • x := semimodule.add_smul r s
7171
variables (R)
7272
@[simp] theorem zero_smul : (0 : R) • x = 0 := semimodule.zero_smul x
7373

74+
theorem two_smul : (2 : R) • x = x + x := by rw [bit0, add_smul, one_smul]
75+
7476
variable (M)
7577

7678
/-- `(•)` as an `add_monoid_hom`. -/

src/analysis/normed_space/basic.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -633,6 +633,8 @@ hf.div hg hnz
633633

634634
lemma real.norm_eq_abs (r : ℝ) : norm r = abs r := rfl
635635

636+
@[simp] lemma real.norm_two : ∥(2:ℝ)∥ = 2 := abs_of_pos (@two_pos ℝ _)
637+
636638
@[simp] lemma norm_norm [normed_group α] (x : α) : ∥∥x∥∥ = ∥x∥ :=
637639
by rw [real.norm_eq_abs, abs_of_nonneg (norm_nonneg _)]
638640

Lines changed: 98 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,98 @@
1+
/-
2+
Copyright (c) 2020 Yury Kudryashov. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Author: Yury Kudryashov
5+
-/
6+
import algebra.midpoint
7+
import topology.metric_space.isometry
8+
9+
/-!
10+
# Reflection in a point as an `isometric` homeomorphism
11+
12+
Given a `normed_group E` and `x : E`, we define `isometric.reflection x` to be
13+
the reflection in `x` interpreted as an `isometric` homeomorphism.
14+
15+
Reflection is defined as an `equiv.perm` in `data.equiv.mul_add`. In this file
16+
we restate results about `equiv.reflection` for an `isometric.reflection`, and
17+
add a few results about `dist`.
18+
19+
## Tags
20+
21+
reflection, isometric
22+
-/
23+
24+
variables (R : Type*) {E : Type*}
25+
26+
lemma equiv.reflection_fixed_iff_of_module [ring R] [invertible (2:R)]
27+
[add_comm_group E] [module R E] {x y : E} :
28+
(equiv.reflection x : E → E) y = y ↔ y = x :=
29+
equiv.reflection_fixed_iff_of_bit0_inj $ λ x y h,
30+
by rw [← one_smul R x, ← one_smul R y, ← inv_of_mul_self (2:R), mul_smul, mul_smul, two_smul,
31+
two_smul, ← bit0, ← bit0, h]
32+
33+
namespace isometric
34+
35+
section normed_group
36+
37+
variables [normed_group E]
38+
39+
/-- Reflection in `x` as an `isometric` homeomorphism. -/
40+
def reflection (x : E) : E ≃ᵢ E :=
41+
(isometric.neg E).trans (isometric.add_left (x + x))
42+
43+
lemma reflection_apply (x y : E) : (reflection x : E → E) y = x + x - y := rfl
44+
45+
@[simp] lemma reflection_to_equiv (x : E) : (reflection x).to_equiv = equiv.reflection x := rfl
46+
47+
@[simp] lemma reflection_self (x : E) : (reflection x : E → E) x = x := add_sub_cancel _ _
48+
49+
lemma reflection_involutive (x : E) : function.involutive (reflection x : E → E) :=
50+
equiv.reflection_involutive x
51+
52+
@[simp] lemma reflection_symm (x : E) : (reflection x).symm = reflection x :=
53+
to_equiv_inj $ equiv.reflection_symm x
54+
55+
@[simp] lemma reflection_dist_fixed (x y : E) :
56+
dist ((reflection x : E → E) y) x = dist y x :=
57+
by rw [← (reflection x).dist_eq y x, reflection_self]
58+
59+
lemma reflection_dist_self' (x y : E) :
60+
dist ((reflection x : E → E) y) y = dist (x + x) (y + y) :=
61+
by { simp only [reflection_apply, dist_eq_norm], congr' 1, abel }
62+
63+
end normed_group
64+
65+
section module
66+
67+
variables [ring R] [invertible (2:R)] [normed_group E] [module R E]
68+
69+
@[simp] lemma reflection_midpoint_left (x y : E) : (reflection (midpoint R x y) : E → E) x = y :=
70+
equiv.reflection_midpoint_left R x y
71+
72+
@[simp] lemma reflection_midpoint_right (x y : E) : (reflection (midpoint R x y) : E → E) y = x :=
73+
equiv.reflection_midpoint_right R x y
74+
75+
variable (R)
76+
77+
include R
78+
79+
lemma reflection_fixed_iff {x y : E} : (reflection x : E → E) y = y ↔ y = x :=
80+
equiv.reflection_fixed_iff_of_module R
81+
82+
end module
83+
84+
section normed_space
85+
86+
variables (R) [normed_field R] [normed_group E] [normed_space R E]
87+
88+
lemma reflection_dist_self (x y : E) :
89+
dist ((reflection x : E → E) y) y = ∥(2:R)∥ * dist x y :=
90+
by simp only [reflection_dist_self', ← two_smul R x, ← two_smul R y, dist_smul]
91+
92+
end normed_space
93+
94+
lemma reflection_dist_self_real [normed_group E] [normed_space ℝ E] (x y : E) :
95+
dist ((reflection x : E → E) y) y = 2 * dist x y :=
96+
by simp [reflection_dist_self ℝ x y, real.norm_two]
97+
98+
end isometric

src/data/equiv/mul_add.lean

Lines changed: 26 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -283,4 +283,30 @@ protected def inv : perm G :=
283283

284284
end group
285285

286+
section reflection
287+
288+
variables [add_comm_group A] (x y : A)
289+
290+
/-- Reflection in `x` as a permutation. -/
291+
def reflection (x : A) : perm A :=
292+
(equiv.neg A).trans (equiv.add_left (x + x))
293+
294+
lemma reflection_apply : reflection x y = x + x - y := rfl
295+
296+
@[simp] lemma reflection_self : reflection x x = x := add_sub_cancel _ _
297+
298+
lemma reflection_involutive : function.involutive (reflection x : A → A) :=
299+
λ y, by simp only [reflection_apply, sub_sub_cancel]
300+
301+
@[simp] lemma reflection_symm : (reflection x).symm = reflection x :=
302+
by { ext y, rw [symm_apply_eq, reflection_involutive x y] }
303+
304+
/-- `x` is the only fixed point of `reflection x`. This lemma requires `x + x = y + y ↔ x = y`.
305+
There is no typeclass to use here, so we add it as an explicit argument. -/
306+
lemma reflection_fixed_iff_of_bit0_inj {x y : A} (h : function.injective (bit0 : A → A)) :
307+
reflection x y = y ↔ y = x :=
308+
sub_eq_iff_eq_add.trans $ h.eq_iff.trans eq_comm
309+
310+
end reflection
311+
286312
end equiv

src/topology/metric_space/gromov_hausdorff.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -275,10 +275,10 @@ begin
275275
repeat {split},
276276
{ exact λx y, calc
277277
F (inl x, inl y) = dist (Φ x) (Φ y) : rfl
278-
... = dist x y : Φisom.dist_eq },
278+
... = dist x y : Φisom.dist_eq x y },
279279
{ exact λx y, calc
280280
F (inr x, inr y) = dist (Ψ x) (Ψ y) : rfl
281-
... = dist x y : Ψisom.dist_eq },
281+
... = dist x y : Ψisom.dist_eq x y },
282282
{ exact λx y, dist_comm _ _ },
283283
{ exact λx y z, dist_triangle _ _ _ },
284284
{ exact λx y, calc

0 commit comments

Comments
 (0)