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

Commit 10f6c15

Browse files
committed
chore(analysis/normed_space/inner_product): notation for orthogonal complement (#5536)
Notation for `submodule.orthogonal`, so that one can write `Kᗮ` instead of `K.orthogonal`. Simultaneous PR leanprover/vscode-lean#246 adds `\perp` as vscode shorthand for this symbol. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/New.20linear.20algebra.20notation)
1 parent cca6365 commit 10f6c15

File tree

4 files changed

+74
-77
lines changed

4 files changed

+74
-77
lines changed

src/analysis/normed_space/dual.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -182,9 +182,9 @@ begin
182182
exact ⟨0, by simp [hℓ]⟩ },
183183
{ have Ycomplete := is_complete_ker ℓ,
184184
rw [submodule.eq_top_iff_orthogonal_eq_bot Ycomplete, ←hY] at htriv,
185-
change Y.orthogonal ≠ ⊥ at htriv,
185+
change Yᗮ ≠ ⊥ at htriv,
186186
rw [submodule.ne_bot_iff] at htriv,
187-
obtain ⟨z : E, hz : z ∈ Y.orthogonal, z_ne_0 : z ≠ 0⟩ := htriv,
187+
obtain ⟨z : E, hz : z ∈ Yᗮ, z_ne_0 : z ≠ 0⟩ := htriv,
188188
refine ⟨((ℓ z)† / ⟪z, z⟫) • z, _⟩,
189189
ext x,
190190
have h₁ : (ℓ z) • x - (ℓ x) • z ∈ Y,

src/analysis/normed_space/inner_product.lean

Lines changed: 42 additions & 45 deletions
Original file line numberDiff line numberDiff line change
@@ -39,6 +39,8 @@ We globally denote the real and complex inner products by `⟪·, ·⟫_ℝ` and
3939
We also provide two notation namespaces: `real_inner_product_space`, `complex_inner_product_space`,
4040
which respectively introduce the plain notation `⟪·, ·⟫` for the the real and complex inner product.
4141
42+
The orthogonal complement of a submodule `K` is denoted by `Kᗮ`.
43+
4244
## Implementation notes
4345
4446
We choose the convention that inner products are conjugate linear in the first argument and linear
@@ -1634,37 +1636,37 @@ def submodule.orthogonal (K : submodule 𝕜 E) : submodule 𝕜 E :=
16341636
add_mem' := λ x y hx hy u hu, by rw [inner_add_right, hx u hu, hy u hu, add_zero],
16351637
smul_mem' := λ c x hx u hu, by rw [inner_smul_right, hx u hu, mul_zero] }
16361638

1637-
/-- When a vector is in `K.orthogonal`. -/
1638-
lemma submodule.mem_orthogonal (K : submodule 𝕜 E) (v : E) :
1639-
v ∈ K.orthogonal ↔ ∀ u ∈ K, ⟪u, v⟫ = 0 :=
1639+
notation K`ᗮ`:1200 := submodule.orthogonal K
1640+
1641+
/-- When a vector is in `Kᗮ`. -/
1642+
lemma submodule.mem_orthogonal (K : submodule 𝕜 E) (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪u, v⟫ = 0 :=
16401643
iff.rfl
16411644

1642-
/-- When a vector is in `K.orthogonal`, with the inner product the
1645+
/-- When a vector is in `Kᗮ`, with the inner product the
16431646
other way round. -/
1644-
lemma submodule.mem_orthogonal' (K : submodule 𝕜 E) (v : E) :
1645-
v ∈ K.orthogonal ↔ ∀ u ∈ K, ⟪v, u⟫ = 0 :=
1647+
lemma submodule.mem_orthogonal' (K : submodule 𝕜 E) (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪v, u⟫ = 0 :=
16461648
by simp_rw [submodule.mem_orthogonal, inner_eq_zero_sym]
16471649

1648-
/-- A vector in `K` is orthogonal to one in `K.orthogonal`. -/
1650+
/-- A vector in `K` is orthogonal to one in `Kᗮ`. -/
16491651
lemma submodule.inner_right_of_mem_orthogonal {u v : E} {K : submodule 𝕜 E} (hu : u ∈ K)
1650-
(hv : v ∈ K.orthogonal) : ⟪u, v⟫ = 0 :=
1652+
(hv : v ∈ Kᗮ) : ⟪u, v⟫ = 0 :=
16511653
(K.mem_orthogonal v).1 hv u hu
16521654

1653-
/-- A vector in `K.orthogonal` is orthogonal to one in `K`. -/
1655+
/-- A vector in `Kᗮ` is orthogonal to one in `K`. -/
16541656
lemma submodule.inner_left_of_mem_orthogonal {u v : E} {K : submodule 𝕜 E} (hu : u ∈ K)
1655-
(hv : v ∈ K.orthogonal) : ⟪v, u⟫ = 0 :=
1657+
(hv : v ∈ Kᗮ) : ⟪v, u⟫ = 0 :=
16561658
by rw [inner_eq_zero_sym]; exact submodule.inner_right_of_mem_orthogonal hu hv
16571659

1658-
/-- `K` and `K.orthogonal` have trivial intersection. -/
1659-
lemma submodule.orthogonal_disjoint (K : submodule 𝕜 E) : disjoint K K.orthogonal :=
1660+
/-- `K` and `Kᗮ` have trivial intersection. -/
1661+
lemma submodule.orthogonal_disjoint (K : submodule 𝕜 E) : disjoint K Kᗮ :=
16601662
begin
16611663
simp_rw [submodule.disjoint_def, submodule.mem_orthogonal],
16621664
exact λ x hx ho, inner_self_eq_zero.1 (ho x hx)
16631665
end
16641666

1665-
/-- `K.orthogonal` can be characterized as the intersection of the kernels of the operations of
1667+
/-- `Kᗮ` can be characterized as the intersection of the kernels of the operations of
16661668
inner product with each of the elements of `K`. -/
1667-
lemma orthogonal_eq_inter (K : submodule 𝕜 E) : K.orthogonal = ⨅ v : K, (inner_right (v:E)).ker :=
1669+
lemma orthogonal_eq_inter (K : submodule 𝕜 E) : Kᗮ = ⨅ v : K, (inner_right (v:E)).ker :=
16681670
begin
16691671
apply le_antisymm,
16701672
{ rw le_infi_iff,
@@ -1676,15 +1678,15 @@ begin
16761678
end
16771679

16781680
/-- The orthogonal complement of any submodule `K` is closed. -/
1679-
lemma submodule.is_closed_orthogonal (K : submodule 𝕜 E) : is_closed (K.orthogonal : set E) :=
1681+
lemma submodule.is_closed_orthogonal (K : submodule 𝕜 E) : is_closed (Kᗮ : set E) :=
16801682
begin
16811683
rw orthogonal_eq_inter K,
16821684
convert is_closed_Inter (λ v : K, (inner_right (v:E)).is_closed_ker),
16831685
simp
16841686
end
16851687

16861688
/-- In a complete space, the orthogonal complement of any submodule `K` is complete. -/
1687-
instance [complete_space E] (K : submodule 𝕜 E) : complete_space K.orthogonal :=
1689+
instance [complete_space E] (K : submodule 𝕜 E) : complete_space Kᗮ :=
16881690
K.is_closed_orthogonal.complete_space_coe
16891691

16901692
variables (𝕜 E)
@@ -1701,35 +1703,31 @@ variables {𝕜 E}
17011703

17021704
/-- `submodule.orthogonal` reverses the `≤` ordering of two
17031705
subspaces. -/
1704-
lemma submodule.orthogonal_le {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂) :
1705-
K₂.orthogonal ≤ K₁.orthogonal :=
1706+
lemma submodule.orthogonal_le {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂) : K₂ᗮ ≤ K₁ᗮ :=
17061707
(submodule.orthogonal_gc 𝕜 E).monotone_l h
17071708

1708-
/-- `K` is contained in `K.orthogonal.orthogonal`. -/
1709-
lemma submodule.le_orthogonal_orthogonal (K : submodule 𝕜 E) : K ≤ K.orthogonal.orthogonal :=
1709+
/-- `K` is contained in `Kᗮᗮ`. -/
1710+
lemma submodule.le_orthogonal_orthogonal (K : submodule 𝕜 E) : K ≤ Kᗮᗮ :=
17101711
(submodule.orthogonal_gc 𝕜 E).le_u_l _
17111712

17121713
/-- The inf of two orthogonal subspaces equals the subspace orthogonal
17131714
to the sup. -/
1714-
lemma submodule.inf_orthogonal (K₁ K₂ : submodule 𝕜 E) :
1715-
K₁.orthogonal ⊓ K₂.orthogonal = (K₁ ⊔ K₂).orthogonal :=
1715+
lemma submodule.inf_orthogonal (K₁ K₂ : submodule 𝕜 E) : K₁ᗮ ⊓ K₂ᗮ = (K₁ ⊔ K₂)ᗮ :=
17161716
(submodule.orthogonal_gc 𝕜 E).l_sup.symm
17171717

17181718
/-- The inf of an indexed family of orthogonal subspaces equals the
17191719
subspace orthogonal to the sup. -/
1720-
lemma submodule.infi_orthogonal {ι : Type*} (K : ι → submodule 𝕜 E) :
1721-
(⨅ i, (K i).orthogonal) = (supr K).orthogonal :=
1720+
lemma submodule.infi_orthogonal {ι : Type*} (K : ι → submodule 𝕜 E) : (⨅ i, (K i)ᗮ) = (supr K)ᗮ :=
17221721
(submodule.orthogonal_gc 𝕜 E).l_supr.symm
17231722

17241723
/-- The inf of a set of orthogonal subspaces equals the subspace
17251724
orthogonal to the sup. -/
1726-
lemma submodule.Inf_orthogonal (s : set $ submodule 𝕜 E) :
1727-
(⨅ K ∈ s, submodule.orthogonal K) = (Sup s).orthogonal :=
1725+
lemma submodule.Inf_orthogonal (s : set $ submodule 𝕜 E) : (⨅ K ∈ s, Kᗮ) = (Sup s)ᗮ :=
17281726
(submodule.orthogonal_gc 𝕜 E).l_Sup.symm
17291727

1730-
/-- If `K₁` is complete and contained in `K₂`, `K₁` and `K₁.orthogonal ⊓ K₂` span `K₂`. -/
1728+
/-- If `K₁` is complete and contained in `K₂`, `K₁` and `K₁ ⊓ K₂` span `K₂`. -/
17311729
lemma submodule.sup_orthogonal_inf_of_is_complete {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂)
1732-
(hc : is_complete (K₁ : set E)) : K₁ ⊔ (K₁.orthogonal ⊓ K₂) = K₂ :=
1730+
(hc : is_complete (K₁ : set E)) : K₁ ⊔ (K₁ ⊓ K₂) = K₂ :=
17331731
begin
17341732
ext x,
17351733
rw submodule.mem_sup,
@@ -1742,34 +1740,33 @@ begin
17421740
add_sub_cancel'_right _ _⟩ }
17431741
end
17441742

1745-
/-- If `K` is complete, `K` and `K.orthogonal` span the whole
1743+
/-- If `K` is complete, `K` and `Kᗮ` span the whole
17461744
space. -/
17471745
lemma submodule.sup_orthogonal_of_is_complete {K : submodule 𝕜 E} (h : is_complete (K : set E)) :
1748-
K ⊔ K.orthogonal = ⊤ :=
1746+
K ⊔ Kᗮ = ⊤ :=
17491747
begin
17501748
convert submodule.sup_orthogonal_inf_of_is_complete (le_top : K ≤ ⊤) h,
17511749
simp
17521750
end
17531751

1754-
/-- If `K` is complete, `K` and `K.orthogonal` span the whole space. Version using `complete_space`.
1752+
/-- If `K` is complete, `K` and `Kᗮ` span the whole space. Version using `complete_space`.
17551753
-/
17561754
lemma submodule.sup_orthogonal_of_complete_space {K : submodule 𝕜 E} [complete_space K] :
1757-
K ⊔ K.orthogonal = ⊤ :=
1755+
K ⊔ Kᗮ = ⊤ :=
17581756
submodule.sup_orthogonal_of_is_complete (complete_space_coe_iff_is_complete.mp ‹_›)
17591757

17601758
/-- If `K` is complete, any `v` in `E` can be expressed as a sum of elements of `K` and
1761-
`K.orthogonal`. -/
1759+
`Kᗮ`. -/
17621760
lemma submodule.exists_sum_mem_mem_orthogonal (K : submodule 𝕜 E) [complete_space K] (v : E) :
1763-
∃ (y ∈ K) (z ∈ K.orthogonal), v = y + z :=
1761+
∃ (y ∈ K) (z ∈ Kᗮ), v = y + z :=
17641762
begin
1765-
have h_mem : v ∈ K ⊔ K.orthogonal := by simp [submodule.sup_orthogonal_of_complete_space],
1763+
have h_mem : v ∈ K ⊔ Kᗮ := by simp [submodule.sup_orthogonal_of_complete_space],
17661764
obtain ⟨y, hy, z, hz, hyz⟩ := submodule.mem_sup.mp h_mem,
17671765
exact ⟨y, hy, z, hz, hyz.symm⟩
17681766
end
17691767

17701768
/-- If `K` is complete, then the orthogonal complement of its orthogonal complement is itself. -/
1771-
@[simp] lemma submodule.orthogonal_orthogonal (K : submodule 𝕜 E) [complete_space K] :
1772-
K.orthogonal.orthogonal = K :=
1769+
@[simp] lemma submodule.orthogonal_orthogonal (K : submodule 𝕜 E) [complete_space K] : Kᗮᗮ = K :=
17731770
begin
17741771
ext v,
17751772
split,
@@ -1784,31 +1781,31 @@ begin
17841781
exact hw v hv }
17851782
end
17861783

1787-
/-- If `K` is complete, `K` and `K.orthogonal` are complements of each
1784+
/-- If `K` is complete, `K` and `Kᗮ` are complements of each
17881785
other. -/
17891786
lemma submodule.is_compl_orthogonal_of_is_complete {K : submodule 𝕜 E}
1790-
(h : is_complete (K : set E)) : is_compl K K.orthogonal :=
1787+
(h : is_complete (K : set E)) : is_compl K Kᗮ :=
17911788
⟨K.orthogonal_disjoint, le_of_eq (submodule.sup_orthogonal_of_is_complete h).symm⟩
17921789

1793-
@[simp] lemma submodule.top_orthogonal_eq_bot : (⊤ : submodule 𝕜 E).orthogonal = ⊥ :=
1790+
@[simp] lemma submodule.top_orthogonal_eq_bot : (⊤ : submodule 𝕜 E) = ⊥ :=
17941791
begin
17951792
ext,
17961793
rw [submodule.mem_bot, submodule.mem_orthogonal],
17971794
exact ⟨λ h, inner_self_eq_zero.mp (h x submodule.mem_top), by { rintro rfl, simp }⟩
17981795
end
17991796

1800-
@[simp] lemma submodule.bot_orthogonal_eq_top : (⊥ : submodule 𝕜 E).orthogonal = ⊤ :=
1797+
@[simp] lemma submodule.bot_orthogonal_eq_top : (⊥ : submodule 𝕜 E) = ⊤ :=
18011798
begin
18021799
rw [← submodule.top_orthogonal_eq_bot, eq_top_iff],
18031800
exact submodule.le_orthogonal_orthogonal ⊤
18041801
end
18051802

18061803
lemma submodule.eq_top_iff_orthogonal_eq_bot {K : submodule 𝕜 E} (hK : is_complete (K : set E)) :
1807-
K = ⊤ ↔ K.orthogonal = ⊥ :=
1804+
K = ⊤ ↔ Kᗮ = ⊥ :=
18081805
begin
18091806
refine ⟨by { rintro rfl, exact submodule.top_orthogonal_eq_bot }, _⟩,
18101807
intro h,
1811-
have : K ⊔ K.orthogonal = ⊤ := submodule.sup_orthogonal_of_is_complete hK,
1808+
have : K ⊔ Kᗮ = ⊤ := submodule.sup_orthogonal_of_is_complete hK,
18121809
rwa [h, sup_comm, bot_sup_eq] at this,
18131810
end
18141811

@@ -1819,10 +1816,10 @@ containined in it, the dimensions of `K₁` and the intersection of its
18191816
orthogonal subspace with `K₂` add to that of `K₂`. -/
18201817
lemma submodule.findim_add_inf_findim_orthogonal {K₁ K₂ : submodule 𝕜 E}
18211818
[finite_dimensional 𝕜 K₂] (h : K₁ ≤ K₂) :
1822-
findim 𝕜 K₁ + findim 𝕜 (K₁.orthogonal ⊓ K₂ : submodule 𝕜 E) = findim 𝕜 K₂ :=
1819+
findim 𝕜 K₁ + findim 𝕜 (K₁ ⊓ K₂ : submodule 𝕜 E) = findim 𝕜 K₂ :=
18231820
begin
18241821
haveI := submodule.finite_dimensional_of_le h,
1825-
have hd := submodule.dim_sup_add_dim_inf_eq K₁ (K₁.orthogonal ⊓ K₂),
1822+
have hd := submodule.dim_sup_add_dim_inf_eq K₁ (K₁ ⊓ K₂),
18261823
rw [←inf_assoc, (submodule.orthogonal_disjoint K₁).eq_bot, bot_inf_eq, findim_bot,
18271824
submodule.sup_orthogonal_inf_of_is_complete h
18281825
(submodule.complete_of_finite_dimensional _)] at hd,

src/geometry/euclidean/basic.lean

Lines changed: 18 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -545,10 +545,10 @@ def orthogonal_projection_fn (s : affine_subspace ℝ P) [nonempty s] [complete_
545545
(p : P) : P :=
546546
classical.some $ inter_eq_singleton_of_nonempty_of_is_compl
547547
(nonempty_subtype.mp ‹_›)
548-
(mk'_nonempty p s.direction.orthogonal)
548+
(mk'_nonempty p s.directionᗮ)
549549
begin
550550
convert submodule.is_compl_orthogonal_of_is_complete (complete_space_coe_iff_is_complete.mp ‹_›),
551-
exact direction_mk' p s.direction.orthogonal
551+
exact direction_mk' p s.directionᗮ
552552
end
553553

554554
/-- The intersection of the subspace and the orthogonal subspace
@@ -558,13 +558,13 @@ setting up the bundled version and should not be used once that is
558558
defined. -/
559559
lemma inter_eq_singleton_orthogonal_projection_fn {s : affine_subspace ℝ P} [nonempty s]
560560
[complete_space s.direction] (p : P) :
561-
(s : set P) ∩ (mk' p s.direction.orthogonal) = {orthogonal_projection_fn s p} :=
561+
(s : set P) ∩ (mk' p s.directionᗮ) = {orthogonal_projection_fn s p} :=
562562
classical.some_spec $ inter_eq_singleton_of_nonempty_of_is_compl
563563
(nonempty_subtype.mp ‹_›)
564-
(mk'_nonempty p s.direction.orthogonal)
564+
(mk'_nonempty p s.directionᗮ)
565565
begin
566566
convert submodule.is_compl_orthogonal_of_is_complete (complete_space_coe_iff_is_complete.mp ‹_›),
567-
exact direction_mk' p s.direction.orthogonal
567+
exact direction_mk' p s.directionᗮ
568568
end
569569

570570
/-- The `orthogonal_projection_fn` lies in the given subspace. This
@@ -582,7 +582,7 @@ subspace. This lemma is only intended for use in setting up the
582582
bundled version and should not be used once that is defined. -/
583583
lemma orthogonal_projection_fn_mem_orthogonal {s : affine_subspace ℝ P} [nonempty s]
584584
[complete_space s.direction] (p : P) :
585-
orthogonal_projection_fn s p ∈ mk' p s.direction.orthogonal :=
585+
orthogonal_projection_fn s p ∈ mk' p s.directionᗮ :=
586586
begin
587587
rw [←mem_coe, ←set.singleton_subset_iff, ←inter_eq_singleton_orthogonal_projection_fn],
588588
exact set.inter_subset_right _ _
@@ -594,8 +594,8 @@ use in setting up the bundled version and should not be used once that
594594
is defined. -/
595595
lemma orthogonal_projection_fn_vsub_mem_direction_orthogonal {s : affine_subspace ℝ P} [nonempty s]
596596
[complete_space s.direction] (p : P) :
597-
orthogonal_projection_fn s p -ᵥ p ∈ s.direction.orthogonal :=
598-
direction_mk' p s.direction.orthogonal
597+
orthogonal_projection_fn s p -ᵥ p ∈ s.directionᗮ :=
598+
direction_mk' p s.directionᗮ
599599
vsub_mem_direction (orthogonal_projection_fn_mem_orthogonal p) (self_mem_mk' _ _)
600600

601601
/-- The orthogonal projection of a point onto a nonempty affine
@@ -613,7 +613,7 @@ def orthogonal_projection (s : affine_subspace ℝ P) [nonempty s] [complete_spa
613613
vadd_mem_of_mem_direction (orthogonal_projection s.direction v).2
614614
(orthogonal_projection_fn_mem p),
615615
have ho : ((orthogonal_projection s.direction) v : V) +ᵥ orthogonal_projection_fn s p ∈
616-
mk' (v +ᵥ p) s.direction.orthogonal,
616+
mk' (v +ᵥ p) s.directionᗮ,
617617
{ rw [←vsub_right_mem_direction_iff_mem (self_mem_mk' _ _) _, direction_mk',
618618
vsub_vadd_eq_vsub_sub, vadd_vsub_assoc, add_comm, add_sub_assoc],
619619
refine submodule.add_mem _ (orthogonal_projection_fn_vsub_mem_direction_orthogonal p) _,
@@ -645,7 +645,7 @@ through the given point is the `orthogonal_projection` of that point
645645
onto the subspace. -/
646646
lemma inter_eq_singleton_orthogonal_projection {s : affine_subspace ℝ P} [nonempty s]
647647
[complete_space s.direction] (p : P) :
648-
(s : set P) ∩ (mk' p s.direction.orthogonal) = {orthogonal_projection s p} :=
648+
(s : set P) ∩ (mk' p s.directionᗮ) = {orthogonal_projection s p} :=
649649
begin
650650
rw ←orthogonal_projection_fn_eq,
651651
exact inter_eq_singleton_orthogonal_projection_fn p
@@ -659,7 +659,7 @@ lemma orthogonal_projection_mem {s : affine_subspace ℝ P} [nonempty s] [comple
659659
/-- The `orthogonal_projection` lies in the orthogonal subspace. -/
660660
lemma orthogonal_projection_mem_orthogonal (s : affine_subspace ℝ P) [nonempty s]
661661
[complete_space s.direction] (p : P) :
662-
↑(orthogonal_projection s p) ∈ mk' p s.direction.orthogonal :=
662+
↑(orthogonal_projection s p) ∈ mk' p s.directionᗮ :=
663663
orthogonal_projection_fn_mem_orthogonal p
664664

665665
/-- Subtracting a point in the given subspace from the
@@ -686,7 +686,7 @@ begin
686686
split,
687687
{ exact λ h, h ▸ orthogonal_projection_mem p },
688688
{ intro h,
689-
have hp : p ∈ ((s : set P) ∩ mk' p s.direction.orthogonal) := ⟨h, self_mem_mk' p _⟩,
689+
have hp : p ∈ ((s : set P) ∩ mk' p s.directionᗮ) := ⟨h, self_mem_mk' p _⟩,
690690
rw [inter_eq_singleton_orthogonal_projection p] at hp,
691691
symmetry,
692692
exact hp }
@@ -728,23 +728,23 @@ mt dist_orthogonal_projection_eq_zero_iff.mp hp
728728
in the orthogonal direction. -/
729729
lemma orthogonal_projection_vsub_mem_direction_orthogonal (s : affine_subspace ℝ P) [nonempty s]
730730
[complete_space s.direction] (p : P) :
731-
(orthogonal_projection s p : P) -ᵥ p ∈ s.direction.orthogonal :=
731+
(orthogonal_projection s p : P) -ᵥ p ∈ s.directionᗮ :=
732732
orthogonal_projection_fn_vsub_mem_direction_orthogonal p
733733

734734
/-- Subtracting the `orthogonal_projection` from `p` produces a result
735735
in the orthogonal direction. -/
736736
lemma vsub_orthogonal_projection_mem_direction_orthogonal (s : affine_subspace ℝ P) [nonempty s]
737737
[complete_space s.direction] (p : P) :
738-
p -ᵥ orthogonal_projection s p ∈ s.direction.orthogonal :=
739-
direction_mk' p s.direction.orthogonal
738+
p -ᵥ orthogonal_projection s p ∈ s.directionᗮ :=
739+
direction_mk' p s.directionᗮ
740740
vsub_mem_direction (self_mem_mk' _ _) (orthogonal_projection_mem_orthogonal s p)
741741

742742
/-- Adding a vector to a point in the given subspace, then taking the
743743
orthogonal projection, produces the original point if the vector was
744744
in the orthogonal direction. -/
745745
lemma orthogonal_projection_vadd_eq_self {s : affine_subspace ℝ P} [nonempty s]
746746
[complete_space s.direction] {p : P} (hp : p ∈ s) {v : V}
747-
(hv : v ∈ s.direction.orthogonal) :
747+
(hv : v ∈ s.directionᗮ) :
748748
orthogonal_projection s (v +ᵥ p) = ⟨p, hp⟩ :=
749749
begin
750750
have h := vsub_orthogonal_projection_mem_direction_orthogonal s (v +ᵥ p),
@@ -788,7 +788,7 @@ adding multiples of the same orthogonal vector to points in the same
788788
subspace. -/
789789
lemma dist_square_smul_orthogonal_vadd_smul_orthogonal_vadd {s : affine_subspace ℝ P}
790790
{p1 p2 : P} (hp1 : p1 ∈ s) (hp2 : p2 ∈ s) (r1 r2 : ℝ) {v : V}
791-
(hv : v ∈ s.direction.orthogonal) :
791+
(hv : v ∈ s.directionᗮ) :
792792
dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) * dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) =
793793
dist p1 p2 * dist p1 p2 + (r1 - r2) * (r1 - r2) * (∥v∥ * ∥v∥) :=
794794
calc dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) * dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2)
@@ -944,7 +944,7 @@ end
944944
produces the negation of that vector plus the point. -/
945945
lemma reflection_orthogonal_vadd {s : affine_subspace ℝ P} [nonempty s]
946946
[complete_space s.direction] {p : P} (hp : p ∈ s) {v : V}
947-
(hv : v ∈ s.direction.orthogonal) : reflection s (v +ᵥ p) = -v +ᵥ p :=
947+
(hv : v ∈ s.directionᗮ) : reflection s (v +ᵥ p) = -v +ᵥ p :=
948948
begin
949949
rw [reflection_apply, orthogonal_projection_vadd_eq_self hp hv, vsub_vadd_eq_vsub_sub],
950950
simp

0 commit comments

Comments
 (0)