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

Commit 2fb109f

Browse files
committed
feat(analysis/inner_product/orientation, geometry/euclidean/oriented_angle): use bundled orthonormal bases (#16475)
Bundled orthonormal bases (as opposed to bases with a mixin predicate `orthonormal`) were defined in #12060, and some of the use cases were switched over in #12253. This PR completes the job, switching to bundled orthonormal bases in `inner_product/orientation` and `euclidean/oriented_angle` as well as in one remaining construction (the standard ` ℝ`-orthonormal basis of `ℂ`) in `inner_product/pi_L2`. Formalized as part of the Sphere Eversion project. The part that I will be using in future PRs is the bundled version of the construction `orthonormal_basis.adjust_to_orientation` in `inner_product/orientation`.
1 parent 73afba4 commit 2fb109f

File tree

3 files changed

+340
-314
lines changed

3 files changed

+340
-314
lines changed

src/analysis/inner_product_space/orientation.lean

Lines changed: 43 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -21,42 +21,64 @@ orientation.
2121
noncomputable theory
2222

2323
variables {E : Type*} [inner_product_space ℝ E]
24-
variables {ι : Type*} [fintype ι] [decidable_eq ι]
2524

2625
open finite_dimensional
2726

28-
/-- `basis.adjust_to_orientation`, applied to an orthonormal basis, produces an orthonormal
29-
basis. -/
30-
lemma orthonormal.orthonormal_adjust_to_orientation [nonempty ι] {e : basis ι ℝ E}
31-
(h : orthonormal ℝ e) (x : orientation ℝ E ι) : orthonormal ℝ (e.adjust_to_orientation x) :=
32-
h.orthonormal_of_forall_eq_or_eq_neg (e.adjust_to_orientation_apply_eq_or_eq_neg x)
27+
section adjust_to_orientation
28+
variables {ι : Type*} [fintype ι] [decidable_eq ι] [nonempty ι] (e : orthonormal_basis ι ℝ E)
29+
(x : orientation ℝ E ι)
3330

34-
/-- An orthonormal basis, indexed by `fin n`, with the given orientation. -/
35-
protected def orientation.fin_orthonormal_basis {n : ℕ} (hn : 0 < n) (h : finrank ℝ E = n)
36-
(x : orientation ℝ E (fin n)) : basis (fin n) ℝ E :=
31+
/-- `orthonormal_basis.adjust_to_orientation`, applied to an orthonormal basis, produces an
32+
orthonormal basis. -/
33+
lemma orthonormal_basis.orthonormal_adjust_to_orientation :
34+
orthonormal ℝ (e.to_basis.adjust_to_orientation x) :=
3735
begin
38-
haveI := fin.pos_iff_nonempty.1 hn,
39-
haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E),
40-
exact (fin_std_orthonormal_basis h).to_basis.adjust_to_orientation x
36+
apply e.orthonormal.orthonormal_of_forall_eq_or_eq_neg,
37+
simpa using e.to_basis.adjust_to_orientation_apply_eq_or_eq_neg x
4138
end
4239

43-
/-- `orientation.fin_orthonormal_basis` is orthonormal. -/
44-
protected lemma orientation.fin_orthonormal_basis_orthonormal {n : ℕ} (hn : 0 < n)
45-
(h : finrank ℝ E = n) (x : orientation ℝ E (fin n)) :
46-
orthonormal ℝ (x.fin_orthonormal_basis hn h) :=
40+
/-- Given an orthonormal basis and an orientation, return an orthonormal basis giving that
41+
orientation: either the original basis, or one constructed by negating a single (arbitrary) basis
42+
vector. -/
43+
def orthonormal_basis.adjust_to_orientation : orthonormal_basis ι ℝ E :=
44+
(e.to_basis.adjust_to_orientation x).to_orthonormal_basis (e.orthonormal_adjust_to_orientation x)
45+
46+
lemma orthonormal_basis.to_basis_adjust_to_orientation :
47+
(e.adjust_to_orientation x).to_basis = e.to_basis.adjust_to_orientation x :=
48+
(e.to_basis.adjust_to_orientation x).to_basis_to_orthonormal_basis _
49+
50+
/-- `adjust_to_orientation` gives an orthonormal basis with the required orientation. -/
51+
@[simp] lemma orthonormal_basis.orientation_adjust_to_orientation :
52+
(e.adjust_to_orientation x).to_basis.orientation = x :=
53+
begin
54+
rw e.to_basis_adjust_to_orientation,
55+
exact e.to_basis.orientation_adjust_to_orientation x,
56+
end
57+
58+
/-- Every basis vector from `adjust_to_orientation` is either that from the original basis or its
59+
negation. -/
60+
lemma orthonormal_basis.adjust_to_orientation_apply_eq_or_eq_neg (i : ι) :
61+
e.adjust_to_orientation x i = e i ∨ e.adjust_to_orientation x i = -(e i) :=
62+
by simpa [← e.to_basis_adjust_to_orientation]
63+
using e.to_basis.adjust_to_orientation_apply_eq_or_eq_neg x i
64+
65+
end adjust_to_orientation
66+
67+
/-- An orthonormal basis, indexed by `fin n`, with the given orientation. -/
68+
protected def orientation.fin_orthonormal_basis {n : ℕ} (hn : 0 < n) (h : finrank ℝ E = n)
69+
(x : orientation ℝ E (fin n)) : orthonormal_basis (fin n) ℝ E :=
4770
begin
4871
haveI := fin.pos_iff_nonempty.1 hn,
4972
haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E),
50-
exact (show orthonormal ℝ (fin_std_orthonormal_basis h).to_basis, -- Note sure how to format this
51-
by simp only [orthonormal_basis.coe_to_basis, orthonormal_basis.orthonormal]
52-
).orthonormal_adjust_to_orientation _
73+
exact (fin_std_orthonormal_basis h).adjust_to_orientation x
5374
end
5475

5576
/-- `orientation.fin_orthonormal_basis` gives a basis with the required orientation. -/
5677
@[simp] lemma orientation.fin_orthonormal_basis_orientation {n : ℕ} (hn : 0 < n)
5778
(h : finrank ℝ E = n) (x : orientation ℝ E (fin n)) :
58-
(x.fin_orthonormal_basis hn h).orientation = x :=
79+
(x.fin_orthonormal_basis hn h).to_basis.orientation = x :=
5980
begin
6081
haveI := fin.pos_iff_nonempty.1 hn,
61-
exact basis.orientation_adjust_to_orientation _ _
82+
haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E),
83+
exact (fin_std_orthonormal_basis h).orientation_adjust_to_orientation x
6284
end

src/analysis/inner_product_space/pi_L2.lean

Lines changed: 32 additions & 41 deletions
Original file line numberDiff line numberDiff line change
@@ -348,6 +348,11 @@ protected def map {G : Type*} [inner_product_space 𝕜 G] (b : orthonormal_basi
348348
(b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) (i : ι) :
349349
b.map L i = L (b i) := rfl
350350

351+
@[simp] protected lemma to_basis_map {G : Type*} [inner_product_space 𝕜 G]
352+
(b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) :
353+
(b.map L).to_basis = b.to_basis.map L.to_linear_equiv :=
354+
rfl
355+
351356
/-- A basis that is orthonormal is an orthonormal basis. -/
352357
def _root_.basis.to_orthonormal_basis (v : basis ι 𝕜 E) (hv : orthonormal 𝕜 v) :
353358
orthonormal_basis ι 𝕜 E :=
@@ -470,66 +475,52 @@ by { classical,
470475

471476
end orthonormal_basis
472477

473-
/-- If `f : E ≃ₗᵢ[𝕜] E'` is a linear isometry of inner product spaces then an orthonormal basis `v`
474-
of `E` determines a linear isometry `e : E' ≃ₗᵢ[𝕜] euclidean_space 𝕜 ι`. This result states that
475-
`e` may be obtained either by transporting `v` to `E'` or by composing with the linear isometry
476-
`E ≃ₗᵢ[𝕜] euclidean_space 𝕜 ι` provided by `v`. -/
477-
@[simp] lemma basis.map_isometry_euclidean_of_orthonormal (v : basis ι 𝕜 E) (hv : orthonormal 𝕜 v)
478-
(f : E ≃ₗᵢ[𝕜] E') :
479-
((v.map f.to_linear_equiv).to_orthonormal_basis (hv.map_linear_isometry_equiv f)).repr =
480-
f.symm.trans (v.to_orthonormal_basis hv).repr :=
481-
linear_isometry_equiv.to_linear_equiv_injective $ v.map_equiv_fun _
482-
483-
/-- `ℂ` is isometric to `ℝ²` with the Euclidean inner product. -/
484-
def complex.isometry_euclidean : ℂ ≃ₗᵢ[ℝ] (euclidean_space ℝ (fin 2)) :=
478+
/-- `![1, I]` is an orthonormal basis for `ℂ` considered as a real inner product space. -/
479+
def complex.orthonormal_basis_one_I : orthonormal_basis (fin 2) ℝ ℂ :=
485480
(complex.basis_one_I.to_orthonormal_basis
486481
begin
487482
rw orthonormal_iff_ite,
488483
intros i, fin_cases i;
489484
intros j; fin_cases j;
490485
simp [real_inner_eq_re_inner]
491-
end).repr
486+
end)
492487

493-
@[simp] lemma complex.isometry_euclidean_symm_apply (x : euclidean_space ℝ (fin 2)) :
494-
complex.isometry_euclidean.symm x = (x 0) + (x 1) * I :=
495-
begin
496-
convert complex.basis_one_I.equiv_fun_symm_apply x,
497-
{ simpa },
498-
{ simp },
499-
end
488+
@[simp] lemma complex.orthonormal_basis_one_I_repr_apply (z : ℂ) :
489+
complex.orthonormal_basis_one_I.repr z = ![z.re, z.im] :=
490+
rfl
500491

501-
lemma complex.isometry_euclidean_proj_eq_self (z : ℂ) :
502-
↑(complex.isometry_euclidean z 0) + ↑(complex.isometry_euclidean z 1) * (I : ℂ) = z :=
503-
by rw [← complex.isometry_euclidean_symm_apply (complex.isometry_euclidean z),
504-
complex.isometry_euclidean.symm_apply_apply z]
492+
@[simp] lemma complex.orthonormal_basis_one_I_repr_symm_apply (x : euclidean_space ℝ (fin 2)) :
493+
complex.orthonormal_basis_one_I.repr.symm x = (x 0) + (x 1) * I :=
494+
rfl
505495

506-
@[simp] lemma complex.isometry_euclidean_apply_zero (z : ℂ) :
507-
complex.isometry_euclidean z 0 = z.re :=
508-
by { conv_rhs { rw ← complex.isometry_euclidean_proj_eq_self z }, simp }
496+
@[simp] lemma complex.to_basis_orthonormal_basis_one_I :
497+
complex.orthonormal_basis_one_I.to_basis = complex.basis_one_I :=
498+
basis.to_basis_to_orthonormal_basis _ _
509499

510-
@[simp] lemma complex.isometry_euclidean_apply_one (z : ℂ) :
511-
complex.isometry_euclidean z 1 = z.im :=
512-
by { conv_rhs { rw ← complex.isometry_euclidean_proj_eq_self z }, simp }
500+
@[simp] lemma complex.coe_orthonormal_basis_one_I :
501+
(complex.orthonormal_basis_one_I : (fin 2) → ℂ) = ![1, I] :=
502+
by simp [complex.orthonormal_basis_one_I]
513503

514504
/-- The isometry between `ℂ` and a two-dimensional real inner product space given by a basis. -/
515-
def complex.isometry_of_orthonormal {v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) : ℂ ≃ₗᵢ[ℝ] F :=
516-
complex.isometry_euclidean.trans (v.to_orthonormal_basis hv).repr.symm
505+
def complex.isometry_of_orthonormal (v : orthonormal_basis (fin 2) ℝ F) : ℂ ≃ₗᵢ[ℝ] F :=
506+
complex.orthonormal_basis_one_I.repr.trans v.repr.symm
517507

518-
@[simp] lemma complex.map_isometry_of_orthonormal {v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v)
508+
@[simp] lemma complex.map_isometry_of_orthonormal (v : orthonormal_basis (fin 2) ℝ F)
519509
(f : F ≃ₗᵢ[ℝ] F') :
520-
complex.isometry_of_orthonormal (hv.map_linear_isometry_equiv f) =
521-
(complex.isometry_of_orthonormal hv).trans f :=
522-
by simp [complex.isometry_of_orthonormal, linear_isometry_equiv.trans_assoc]
510+
complex.isometry_of_orthonormal (v.map f) =
511+
(complex.isometry_of_orthonormal v).trans f :=
512+
by simp [complex.isometry_of_orthonormal, linear_isometry_equiv.trans_assoc, orthonormal_basis.map]
523513

524514
lemma complex.isometry_of_orthonormal_symm_apply
525-
{v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) (f : F) :
526-
(complex.isometry_of_orthonormal hv).symm f = (v.coord 0 f : ℂ) + (v.coord 1 f : ℂ) * I :=
515+
(v : orthonormal_basis (fin 2) ℝ F) (f : F) :
516+
(complex.isometry_of_orthonormal v).symm f
517+
= (v.to_basis.coord 0 f : ℂ) + (v.to_basis.coord 1 f : ℂ) * I :=
527518
by simp [complex.isometry_of_orthonormal]
528519

529520
lemma complex.isometry_of_orthonormal_apply
530-
{v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) (z : ℂ) :
531-
complex.isometry_of_orthonormal hv z = z.re • v 0 + z.im • v 1 :=
532-
by simp [complex.isometry_of_orthonormal, (dec_trivial : (finset.univ : finset (fin 2)) = {0, 1})]
521+
(v : orthonormal_basis (fin 2) ℝ F) (z : ℂ) :
522+
complex.isometry_of_orthonormal v z = z.re • v 0 + z.im • v 1 :=
523+
by simp [complex.isometry_of_orthonormal, ← v.sum_repr_symm]
533524

534525
open finite_dimensional
535526

0 commit comments

Comments
 (0)