Skip to content

Commit 6e974c7

Browse files
committed
feat(Geometry/Euclidean/Angle): angles with orthogonal projections (#30522)
Add some lemmas that angles involving orthogonal projections are right angles. The unoriented versions abstract something used multiple times in `Mathlib.Geometry.Euclidean.Angle.Bisector` (which is thus refactored to use the new lemmas); I expect the oriented versions may be of use in future in further oriented angle bisection lemmas building on top of #30477.
1 parent 1835445 commit 6e974c7

File tree

4 files changed

+110
-15
lines changed

4 files changed

+110
-15
lines changed

Mathlib.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -3918,13 +3918,15 @@ import Mathlib.Geometry.Euclidean.Altitude
39183918
import Mathlib.Geometry.Euclidean.Angle.Bisector
39193919
import Mathlib.Geometry.Euclidean.Angle.Oriented.Affine
39203920
import Mathlib.Geometry.Euclidean.Angle.Oriented.Basic
3921+
import Mathlib.Geometry.Euclidean.Angle.Oriented.Projection
39213922
import Mathlib.Geometry.Euclidean.Angle.Oriented.RightAngle
39223923
import Mathlib.Geometry.Euclidean.Angle.Oriented.Rotation
39233924
import Mathlib.Geometry.Euclidean.Angle.Sphere
39243925
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Affine
39253926
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Basic
39263927
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Conformal
39273928
import Mathlib.Geometry.Euclidean.Angle.Unoriented.CrossProduct
3929+
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Projection
39283930
import Mathlib.Geometry.Euclidean.Angle.Unoriented.RightAngle
39293931
import Mathlib.Geometry.Euclidean.Basic
39303932
import Mathlib.Geometry.Euclidean.Circumcenter

Mathlib/Geometry/Euclidean/Angle/Bisector.lean

Lines changed: 5 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Copyright (c) 2025 Joseph Myers. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Joseph Myers
55
-/
6+
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Projection
67
import Mathlib.Geometry.Euclidean.Angle.Unoriented.RightAngle
78
import Mathlib.Geometry.Euclidean.Projection
89

@@ -41,15 +42,12 @@ private lemma dist_orthogonalProjection_eq_iff_angle_eq_aux₁ {p p' : P}
4142
· subst hpp'
4243
exact hp'.2
4344
· by_contra hn
44-
rw [angle_self_of_ne hpp', angle_comm, angle_eq_arcsin_of_angle_eq_pi_div_two,
45+
rw [angle_self_of_ne hpp', angle_comm,
46+
angle_eq_arcsin_of_angle_eq_pi_div_two (angle_self_orthogonalProjection p hp'.2),
4547
Real.zero_eq_arcsin_iff, div_eq_zero_iff] at h
4648
· simp only [dist_eq_zero, hpp', or_false] at h
4749
rw [eq_comm] at h
4850
simp [orthogonalProjection_eq_self_iff, hn] at h
49-
· rw [angle, ← InnerProductGeometry.inner_eq_zero_iff_angle_eq_pi_div_two]
50-
exact Submodule.inner_left_of_mem_orthogonal (K := s₂.direction)
51-
(AffineSubspace.vsub_mem_direction hp'.2 (orthogonalProjection_mem _))
52-
(vsub_orthogonalProjection_mem_direction_orthogonal _ _)
5351
· exact .inl (Ne.symm (orthogonalProjection_eq_self_iff.symm.not.1 hn))
5452

5553
/-- Auxiliary lemma for the degenerate case of `dist_orthogonalProjection_eq_iff_angle_eq` where
@@ -86,10 +84,10 @@ lemma dist_orthogonalProjection_eq_iff_angle_eq {p p' : P} {s₁ s₂ : AffineSu
8684
· exact dist_orthogonalProjection_eq_iff_angle_eq_aux hp' h'
8785
rw [not_or] at h'
8886
rw [angle_comm,
89-
angle_eq_arcsin_of_angle_eq_pi_div_two ?_
87+
angle_eq_arcsin_of_angle_eq_pi_div_two (angle_self_orthogonalProjection p hp'.1)
9088
(.inl (Ne.symm (orthogonalProjection_eq_self_iff.symm.not.1 h'.1))),
9189
angle_comm,
92-
angle_eq_arcsin_of_angle_eq_pi_div_two ?_
90+
angle_eq_arcsin_of_angle_eq_pi_div_two (angle_self_orthogonalProjection p hp'.2)
9391
(.inl (Ne.symm (orthogonalProjection_eq_self_iff.symm.not.1 h'.2)))]
9492
· refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩
9593
· rw [h]
@@ -105,13 +103,5 @@ lemma dist_orthogonalProjection_eq_iff_angle_eq {p p' : P} {s₁ s₂ : AffineSu
105103
exact Metric.infDist_le_dist_of_mem (SetLike.mem_coe.1 hp'.1)
106104
· rw [dist_orthogonalProjection_eq_infDist]
107105
exact Metric.infDist_le_dist_of_mem (SetLike.mem_coe.1 hp'.2)
108-
· rw [angle, ← InnerProductGeometry.inner_eq_zero_iff_angle_eq_pi_div_two]
109-
exact Submodule.inner_left_of_mem_orthogonal (K := s₂.direction)
110-
(AffineSubspace.vsub_mem_direction hp'.2 (orthogonalProjection_mem _))
111-
(vsub_orthogonalProjection_mem_direction_orthogonal _ _)
112-
· rw [angle, ← InnerProductGeometry.inner_eq_zero_iff_angle_eq_pi_div_two]
113-
exact Submodule.inner_left_of_mem_orthogonal (K := s₁.direction)
114-
(AffineSubspace.vsub_mem_direction hp'.1 (orthogonalProjection_mem _))
115-
(vsub_orthogonalProjection_mem_direction_orthogonal _ _)
116106

117107
end EuclideanGeometry
Lines changed: 63 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,63 @@
1+
/-
2+
Copyright (c) 2025 Joseph Myers. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Joseph Myers
5+
-/
6+
import Mathlib.Geometry.Euclidean.Angle.Oriented.Affine
7+
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Projection
8+
9+
/-!
10+
# Oriented angles and orthogonal projection.
11+
12+
This file proves lemmas relating to oriented angles involving orthogonal projections.
13+
14+
-/
15+
16+
17+
namespace EuclideanGeometry
18+
19+
open Module
20+
open scoped Real
21+
22+
variable {V P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P]
23+
variable [NormedAddTorsor V P] [hd2 : Fact (finrank ℝ V = 2)] [Module.Oriented ℝ V (Fin 2)]
24+
25+
lemma oangle_self_orthogonalProjection (p : P) {p' : P} {s : AffineSubspace ℝ P}
26+
[s.direction.HasOrthogonalProjection] (hp : p ∉ s) (h : p' ∈ s)
27+
(hp' : haveI : Nonempty s := ⟨p', h⟩; p' ≠ orthogonalProjection s p) :
28+
haveI : Nonempty s := ⟨p', h⟩
29+
∡ p (orthogonalProjection s p) p' = (π / 2 : ℝ) ∨
30+
∡ p (orthogonalProjection s p) p' = (-π / 2 : ℝ) := by
31+
haveI : Nonempty s := ⟨p', h⟩
32+
have hpne : p ≠ orthogonalProjection s p := Ne.symm (orthogonalProjection_eq_self_iff.not.2 hp)
33+
have ha := oangle_eq_angle_or_eq_neg_angle hpne hp'
34+
rw [angle_self_orthogonalProjection p h] at ha
35+
rwa [neg_div]
36+
37+
lemma oangle_orthogonalProjection_self (p : P) {p' : P} {s : AffineSubspace ℝ P}
38+
[s.direction.HasOrthogonalProjection] (hp : p ∉ s) (h : p' ∈ s)
39+
(hp' : haveI : Nonempty s := ⟨p', h⟩; p' ≠ orthogonalProjection s p) :
40+
haveI : Nonempty s := ⟨p', h⟩
41+
∡ p' (orthogonalProjection s p) p = (π / 2 : ℝ) ∨
42+
∡ p' (orthogonalProjection s p) p = (-π / 2 : ℝ) := by
43+
rw [oangle_rev, neg_eq_iff_eq_neg, neg_eq_iff_eq_neg, or_comm, ← Real.Angle.coe_neg, neg_div,
44+
neg_neg, ← Real.Angle.coe_neg, ← neg_div]
45+
exact oangle_self_orthogonalProjection p hp h hp'
46+
47+
lemma two_zsmul_oangle_self_orthogonalProjection (p : P) {p' : P} {s : AffineSubspace ℝ P}
48+
[s.direction.HasOrthogonalProjection] (hp : p ∉ s) (h : p' ∈ s)
49+
(hp' : haveI : Nonempty s := ⟨p', h⟩; p' ≠ orthogonalProjection s p) :
50+
haveI : Nonempty s := ⟨p', h⟩
51+
(2 : ℤ) • ∡ p (orthogonalProjection s p) p' = π := by
52+
rw [Real.Angle.two_zsmul_eq_pi_iff]
53+
exact oangle_self_orthogonalProjection p hp h hp'
54+
55+
lemma two_zsmul_oangle_orthogonalProjection_self (p : P) {p' : P} {s : AffineSubspace ℝ P}
56+
[s.direction.HasOrthogonalProjection] (hp : p ∉ s) (h : p' ∈ s)
57+
(hp' : haveI : Nonempty s := ⟨p', h⟩; p' ≠ orthogonalProjection s p) :
58+
haveI : Nonempty s := ⟨p', h⟩
59+
(2 : ℤ) • ∡ p' (orthogonalProjection s p) p = π := by
60+
rw [Real.Angle.two_zsmul_eq_pi_iff]
61+
exact oangle_orthogonalProjection_self p hp h hp'
62+
63+
end EuclideanGeometry
Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,40 @@
1+
/-
2+
Copyright (c) 2025 Joseph Myers. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Joseph Myers
5+
-/
6+
import Mathlib.Geometry.Euclidean.Angle.Unoriented.Affine
7+
import Mathlib.Geometry.Euclidean.Projection
8+
9+
/-!
10+
# Angles and orthogonal projection.
11+
12+
This file proves lemmas relating to angles involving orthogonal projections.
13+
14+
-/
15+
16+
17+
namespace EuclideanGeometry
18+
19+
variable {V P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P]
20+
variable [NormedAddTorsor V P]
21+
22+
open scoped Real
23+
24+
@[simp] lemma angle_self_orthogonalProjection (p : P) {p' : P} {s : AffineSubspace ℝ P}
25+
[s.direction.HasOrthogonalProjection] (h : p' ∈ s) :
26+
haveI : Nonempty s := ⟨p', h⟩
27+
∠ p (orthogonalProjection s p) p' = π / 2 := by
28+
haveI : Nonempty s := ⟨p', h⟩
29+
rw [angle, ← InnerProductGeometry.inner_eq_zero_iff_angle_eq_pi_div_two]
30+
exact Submodule.inner_left_of_mem_orthogonal (K := s.direction)
31+
(AffineSubspace.vsub_mem_direction h (orthogonalProjection_mem _))
32+
(vsub_orthogonalProjection_mem_direction_orthogonal _ _)
33+
34+
@[simp] lemma angle_orthogonalProjection_self (p : P) {p' : P} {s : AffineSubspace ℝ P}
35+
[s.direction.HasOrthogonalProjection] (h : p' ∈ s) :
36+
haveI : Nonempty s := ⟨p', h⟩
37+
∠ p' (orthogonalProjection s p) p = π / 2 := by
38+
rw [angle_comm, angle_self_orthogonalProjection p h]
39+
40+
end EuclideanGeometry

0 commit comments

Comments
 (0)