@@ -39,6 +39,8 @@ We globally denote the real and complex inner products by `⟪·, ·⟫_ℝ` and
39
39
We also provide two notation namespaces: `real_inner_product_space`, `complex_inner_product_space`,
40
40
which respectively introduce the plain notation `⟪·, ·⟫` for the the real and complex inner product.
41
41
42
+ The orthogonal complement of a submodule `K` is denoted by `Kᗮ`.
43
+
42
44
## Implementation notes
43
45
44
46
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 :=
1634
1636
add_mem' := λ x y hx hy u hu, by rw [inner_add_right, hx u hu, hy u hu, add_zero],
1635
1637
smul_mem' := λ c x hx u hu, by rw [inner_smul_right, hx u hu, mul_zero] }
1636
1638
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 :=
1640
1643
iff.rfl
1641
1644
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
1643
1646
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 :=
1646
1648
by simp_rw [submodule.mem_orthogonal, inner_eq_zero_sym]
1647
1649
1648
- /-- A vector in `K` is orthogonal to one in `K.orthogonal `. -/
1650
+ /-- A vector in `K` is orthogonal to one in `Kᗮ `. -/
1649
1651
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 :=
1651
1653
(K.mem_orthogonal v).1 hv u hu
1652
1654
1653
- /-- A vector in `K.orthogonal ` is orthogonal to one in `K`. -/
1655
+ /-- A vector in `Kᗮ ` is orthogonal to one in `K`. -/
1654
1656
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 :=
1656
1658
by rw [inner_eq_zero_sym]; exact submodule.inner_right_of_mem_orthogonal hu hv
1657
1659
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ᗮ :=
1660
1662
begin
1661
1663
simp_rw [submodule.disjoint_def, submodule.mem_orthogonal],
1662
1664
exact λ x hx ho, inner_self_eq_zero.1 (ho x hx)
1663
1665
end
1664
1666
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
1666
1668
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 :=
1668
1670
begin
1669
1671
apply le_antisymm,
1670
1672
{ rw le_infi_iff,
@@ -1676,15 +1678,15 @@ begin
1676
1678
end
1677
1679
1678
1680
/-- 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) :=
1680
1682
begin
1681
1683
rw orthogonal_eq_inter K,
1682
1684
convert is_closed_Inter (λ v : K, (inner_right (v:E)).is_closed_ker),
1683
1685
simp
1684
1686
end
1685
1687
1686
1688
/-- 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ᗮ :=
1688
1690
K.is_closed_orthogonal.complete_space_coe
1689
1691
1690
1692
variables (𝕜 E)
@@ -1701,35 +1703,31 @@ variables {𝕜 E}
1701
1703
1702
1704
/-- `submodule.orthogonal` reverses the `≤` ordering of two
1703
1705
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₁ᗮ :=
1706
1707
(submodule.orthogonal_gc 𝕜 E).monotone_l h
1707
1708
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ᗮᗮ :=
1710
1711
(submodule.orthogonal_gc 𝕜 E).le_u_l _
1711
1712
1712
1713
/-- The inf of two orthogonal subspaces equals the subspace orthogonal
1713
1714
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₂)ᗮ :=
1716
1716
(submodule.orthogonal_gc 𝕜 E).l_sup.symm
1717
1717
1718
1718
/-- The inf of an indexed family of orthogonal subspaces equals the
1719
1719
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)ᗮ :=
1722
1721
(submodule.orthogonal_gc 𝕜 E).l_supr.symm
1723
1722
1724
1723
/-- The inf of a set of orthogonal subspaces equals the subspace
1725
1724
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)ᗮ :=
1728
1726
(submodule.orthogonal_gc 𝕜 E).l_Sup.symm
1729
1727
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₂`. -/
1731
1729
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₂ :=
1733
1731
begin
1734
1732
ext x,
1735
1733
rw submodule.mem_sup,
@@ -1742,34 +1740,33 @@ begin
1742
1740
add_sub_cancel'_right _ _⟩ }
1743
1741
end
1744
1742
1745
- /-- If `K` is complete, `K` and `K.orthogonal ` span the whole
1743
+ /-- If `K` is complete, `K` and `Kᗮ ` span the whole
1746
1744
space. -/
1747
1745
lemma submodule.sup_orthogonal_of_is_complete {K : submodule 𝕜 E} (h : is_complete (K : set E)) :
1748
- K ⊔ K.orthogonal = ⊤ :=
1746
+ K ⊔ Kᗮ = ⊤ :=
1749
1747
begin
1750
1748
convert submodule.sup_orthogonal_inf_of_is_complete (le_top : K ≤ ⊤) h,
1751
1749
simp
1752
1750
end
1753
1751
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`.
1755
1753
-/
1756
1754
lemma submodule.sup_orthogonal_of_complete_space {K : submodule 𝕜 E} [complete_space K] :
1757
- K ⊔ K.orthogonal = ⊤ :=
1755
+ K ⊔ Kᗮ = ⊤ :=
1758
1756
submodule.sup_orthogonal_of_is_complete (complete_space_coe_iff_is_complete.mp ‹_›)
1759
1757
1760
1758
/-- If `K` is complete, any `v` in `E` can be expressed as a sum of elements of `K` and
1761
- `K.orthogonal `. -/
1759
+ `Kᗮ `. -/
1762
1760
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 :=
1764
1762
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],
1766
1764
obtain ⟨y, hy, z, hz, hyz⟩ := submodule.mem_sup.mp h_mem,
1767
1765
exact ⟨y, hy, z, hz, hyz.symm⟩
1768
1766
end
1769
1767
1770
1768
/-- 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 :=
1773
1770
begin
1774
1771
ext v,
1775
1772
split,
@@ -1784,31 +1781,31 @@ begin
1784
1781
exact hw v hv }
1785
1782
end
1786
1783
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
1788
1785
other. -/
1789
1786
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ᗮ :=
1791
1788
⟨K.orthogonal_disjoint, le_of_eq (submodule.sup_orthogonal_of_is_complete h).symm⟩
1792
1789
1793
- @[simp] lemma submodule.top_orthogonal_eq_bot : (⊤ : submodule 𝕜 E).orthogonal = ⊥ :=
1790
+ @[simp] lemma submodule.top_orthogonal_eq_bot : (⊤ : submodule 𝕜 E)ᗮ = ⊥ :=
1794
1791
begin
1795
1792
ext,
1796
1793
rw [submodule.mem_bot, submodule.mem_orthogonal],
1797
1794
exact ⟨λ h, inner_self_eq_zero.mp (h x submodule.mem_top), by { rintro rfl, simp }⟩
1798
1795
end
1799
1796
1800
- @[simp] lemma submodule.bot_orthogonal_eq_top : (⊥ : submodule 𝕜 E).orthogonal = ⊤ :=
1797
+ @[simp] lemma submodule.bot_orthogonal_eq_top : (⊥ : submodule 𝕜 E)ᗮ = ⊤ :=
1801
1798
begin
1802
1799
rw [← submodule.top_orthogonal_eq_bot, eq_top_iff],
1803
1800
exact submodule.le_orthogonal_orthogonal ⊤
1804
1801
end
1805
1802
1806
1803
lemma submodule.eq_top_iff_orthogonal_eq_bot {K : submodule 𝕜 E} (hK : is_complete (K : set E)) :
1807
- K = ⊤ ↔ K.orthogonal = ⊥ :=
1804
+ K = ⊤ ↔ Kᗮ = ⊥ :=
1808
1805
begin
1809
1806
refine ⟨by { rintro rfl, exact submodule.top_orthogonal_eq_bot }, _⟩,
1810
1807
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,
1812
1809
rwa [h, sup_comm, bot_sup_eq] at this ,
1813
1810
end
1814
1811
@@ -1819,10 +1816,10 @@ containined in it, the dimensions of `K₁` and the intersection of its
1819
1816
orthogonal subspace with `K₂` add to that of `K₂`. -/
1820
1817
lemma submodule.findim_add_inf_findim_orthogonal {K₁ K₂ : submodule 𝕜 E}
1821
1818
[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₂ :=
1823
1820
begin
1824
1821
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₂),
1826
1823
rw [←inf_assoc, (submodule.orthogonal_disjoint K₁).eq_bot, bot_inf_eq, findim_bot,
1827
1824
submodule.sup_orthogonal_inf_of_is_complete h
1828
1825
(submodule.complete_of_finite_dimensional _)] at hd,
0 commit comments