Skip to content

Conversation

@jsm28
Copy link
Collaborator

@jsm28 jsm28 commented Oct 12, 2025

Add a lemma

lemma dist_orthogonalProjection_eq_iff_oangle_eq {p p' : P} {s₁ s₂ : AffineSubspace ℝ P}
    [s₁.direction.HasOrthogonalProjection] [s₂.direction.HasOrthogonalProjection]
    (hp' : p' ∈ s₁ ⊓ s₂)
    (hne : haveI : Nonempty s₁ := ⟨p', hp'.1⟩; haveI : Nonempty s₂ := ⟨p', hp'.2⟩;
      (orthogonalProjection s₁ p : P) ≠ orthogonalProjection s₂ p)
    (hp₁ : haveI : Nonempty s₁ := ⟨p', hp'.1⟩; orthogonalProjection s₁ p ≠ p')
    (hp₂ : haveI : Nonempty s₂ := ⟨p', hp'.2⟩; orthogonalProjection s₂ p ≠ p') :
    haveI : Nonempty s₁ := ⟨p', hp'.1haveI : Nonempty s₂ := ⟨p', hp'.2⟩
    dist p (orthogonalProjection s₁ p) = dist p (orthogonalProjection s₂ p) ↔
      ∡ (orthogonalProjection s₁ p : P) p' p = ∡ p p' (orthogonalProjection s₂ p) :=

that is an oriented angle analogue of the existing dist_orthogonalProjection_eq_iff_angle_eq. Because the minimal nondegeneracy conditions required for the two directions of this lemma are different (whereas the unoriented version doesn't need any nondegeneracy conditions), those two directions are added as separate lemmas, each with minimal nondegeneracy conditions, from which the iff version is then deduced.


Open in Gitpod

jsm28 added 4 commits October 12, 2025 19:56
…Projection_of_le`

Add a lemma that `orthogonalProjection s₁ (orthogonalProjection s₂ p)
= orthogonalProjection s₁ p` for `s₁ ≤ s₂`. (A similar lemma for
projection onto submodules rather than affine subspaces already
exists.)
… and bisection lemmas

Add lemmas

```lean
lemma oangle_eq_neg_of_angle_eq_of_sign_eq_neg {w x y z : V}
    (h : InnerProductGeometry.angle w x = InnerProductGeometry.angle y z)
    (hs : (o.oangle w x).sign = -(o.oangle y z).sign) : o.oangle w x = -o.oangle y z := by
```

and

```lean
lemma angle_eq_iff_oangle_eq_neg_of_sign_eq_neg {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0)
    (hy : y ≠ 0) (hz : z ≠ 0) (hs : (o.oangle w x).sign = -(o.oangle y z).sign) :
    InnerProductGeometry.angle w x = InnerProductGeometry.angle y z ↔
      o.oangle w x = -o.oangle y z := by
```

and similar affine versions, corresponding to such lemmas that already
exist when the signs are equal rather than negations of each other.
Deduce lemmas relating oriented and unoriented versions of angle
bisection:

```lean
lemma angle_eq_iff_oangle_eq_or_sameRay {x y z : V} (hx : x ≠ 0) (hz : z ≠ 0) :
    InnerProductGeometry.angle x y = InnerProductGeometry.angle y z ↔
      o.oangle x y = o.oangle y z ∨ SameRay ℝ x z := by
```

and

```lean
lemma angle_eq_iff_oangle_eq_or_wbtw {p₁ p₂ p₃ p₄ : P} (hp₁ : p₁ ≠ p₂) (hp₄ : p₄ ≠ p₂) :
    ∠ p₁ p₂ p₃ = ∠ p₃ p₂ p₄ ↔ ∡ p₁ p₂ p₃ = ∡ p₃ p₂ p₄ ∨ Wbtw ℝ p₂ p₁ p₄ ∨ Wbtw ℝ p₂ p₄ p₁ := by
```
… equal distance

Add a lemma

```lean
lemma dist_orthogonalProjection_eq_iff_oangle_eq {p p' : P} {s₁ s₂ : AffineSubspace ℝ P}
    [s₁.direction.HasOrthogonalProjection] [s₂.direction.HasOrthogonalProjection]
    (hp' : p' ∈ s₁ ⊓ s₂)
    (hne : haveI : Nonempty s₁ := ⟨p', hp'.1⟩; haveI : Nonempty s₂ := ⟨p', hp'.2⟩;
      (orthogonalProjection s₁ p : P) ≠ orthogonalProjection s₂ p)
    (hp₁ : haveI : Nonempty s₁ := ⟨p', hp'.1⟩; orthogonalProjection s₁ p ≠ p')
    (hp₂ : haveI : Nonempty s₂ := ⟨p', hp'.2⟩; orthogonalProjection s₂ p ≠ p') :
    haveI : Nonempty s₁ := ⟨p', hp'.1⟩
    haveI : Nonempty s₂ := ⟨p', hp'.2⟩
    dist p (orthogonalProjection s₁ p) = dist p (orthogonalProjection s₂ p) ↔
      ∡ (orthogonalProjection s₁ p : P) p' p = ∡ p p' (orthogonalProjection s₂ p) :=
```

that is an oriented angle analogue of the existing
`dist_orthogonalProjection_eq_iff_angle_eq`.  Because the minimal
nondegeneracy conditions required for the two directions of this lemma
are different (whereas the unoriented version doesn't need any
nondegeneracy conditions), those two directions are added as separate
lemmas, each with minimal nondegeneracy conditions, from which the
`iff` version is then deduced.
@jsm28 jsm28 added the t-euclidean-geometry Affine and axiomatic geometry label Oct 12, 2025
@github-actions github-actions bot added the large-import Automatically added label for PRs with a significant increase in transitive imports label Oct 12, 2025
@github-actions
Copy link

github-actions bot commented Oct 12, 2025

PR summary ed54e09893

Import changes exceeding 2%

% File
+2.18% Mathlib.Geometry.Euclidean.Angle.Bisector

Import changes for modified files

Dependency changes

File Base Count Head Count Change
Mathlib.Geometry.Euclidean.Angle.Bisector 2114 2160 +46 (+2.18%)
Import changes for all files
Files Import difference
Mathlib.Geometry.Euclidean.Angle.Bisector 46

Declarations diff

+ dist_orthogonalProjection_eq_iff_oangle_eq
+ dist_orthogonalProjection_eq_of_oangle_eq
+ oangle_eq_of_dist_orthogonalProjection_eq

You can run this locally as follows
## summary with just the declaration names:
./scripts/declarations_diff.sh <optional_commit>

## more verbose report:
./scripts/declarations_diff.sh long <optional_commit>

The doc-module for script/declarations_diff.sh contains some details about this script.


No changes to technical debt.

You can run this locally as

./scripts/technical-debt-metrics.sh pr_summary
  • The relative value is the weighted sum of the differences with weight given by the inverse of the current value of the statistic.
  • The absolute value is the relative value divided by the total sum of the inverses of the current values (i.e. the weighted average of the differences).

@jsm28
Copy link
Collaborator Author

jsm28 commented Oct 12, 2025

The increased imports are explicitly intended for this file and just a consequence of building up this file over multiple PRs rather than adding oriented and unoriented lemmas together, as noted in #29656 when I added Bisector.lean.

@mathlib4-dependent-issues-bot mathlib4-dependent-issues-bot added the blocked-by-other-PR This PR depends on another PR (this label is automatically managed by a bot) label Oct 12, 2025
jsm28 added a commit to jsm28/mathlib4 that referenced this pull request Oct 14, 2025
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 leanprover-community#30477.
jsm28 and others added 19 commits October 19, 2025 21:39
Add a lemma

```lean
lemma orthogonalProjection_eq_iff_mem {s : AffineSubspace ℝ P} [Nonempty s]
    [s.direction.HasOrthogonalProjection] {p q : P} :
    orthogonalProjection s p = q ↔ q ∈ s ∧ p -ᵥ q ∈ s.directionᗮ := by
```

that gives the characteristic property of the orthogonal projection in
a more convenient form to use than the existing
`inter_eq_singleton_orthogonalProjection` (from which it is derived).
Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
Co-authored-by: Jovan Gerbscheid <56355248+JovanGerb@users.noreply.github.com>
mathlib-bors bot pushed a commit that referenced this pull request Oct 27, 2025
…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.
jsm28 added a commit to jsm28/mathlib4 that referenced this pull request Oct 27, 2025
Add lemmas relating equal distance to two lines to bisecting the angle between them mod π (expressed as usual as equality of twice angles), building on the previous lemmas (leanprover-community#30477) dealing with bisection mod 2π.
@mathlib4-dependent-issues-bot mathlib4-dependent-issues-bot removed the blocked-by-other-PR This PR depends on another PR (this label is automatically managed by a bot) label Oct 28, 2025
Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment

Labels

large-import Automatically added label for PRs with a significant increase in transitive imports t-euclidean-geometry Affine and axiomatic geometry

Projects

None yet

Development

Successfully merging this pull request may close these issues.

2 participants