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

Commit 0179605

Browse files
committed
feat(analysis/inner_product_space/positive): definition and basic facts about positive operators (#15470)
1 parent 983c737 commit 0179605

File tree

2 files changed

+198
-0
lines changed

2 files changed

+198
-0
lines changed

src/analysis/inner_product_space/adjoint.lean

Lines changed: 72 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -77,6 +77,26 @@ by rw [hT x y, inner_conj_sym]
7777
⟪T x, y⟫ = ⟪x, T y⟫ :=
7878
hT x y
7979

80+
lemma is_self_adjoint_zero : is_self_adjoint (0 : E →ₗ[𝕜] E) :=
81+
λ x y, (inner_zero_right : ⟪x, 0⟫ = 0).symm ▸ (inner_zero_left : ⟪0, y⟫ = 0)
82+
83+
lemma is_self_adjoint_id : is_self_adjoint (linear_map.id : E →ₗ[𝕜] E) :=
84+
λ x y, rfl
85+
86+
lemma is_self_adjoint.add {T S : E →ₗ[𝕜] E} (hT : is_self_adjoint T)
87+
(hS : is_self_adjoint S) : is_self_adjoint (T + S) :=
88+
begin
89+
intros x y,
90+
rw [linear_map.add_apply, inner_add_left, hT x y, hS x y, ← inner_add_right],
91+
refl
92+
end
93+
94+
/-- The orthogonal projection is self-adjoint. -/
95+
lemma orthogonal_projection_is_self_adjoint [complete_space E] (U : submodule 𝕜 E)
96+
[complete_space U] :
97+
is_self_adjoint (U.subtypeL ∘L orthogonal_projection U : E →ₗ[𝕜] E):=
98+
inner_orthogonal_projection_left_eq_right U
99+
80100
/-- The **Hellinger--Toeplitz theorem**: if a symmetric operator is defined everywhere, then
81101
it is automatically continuous. -/
82102
lemma is_self_adjoint.continuous [complete_space E] {T : E →ₗ[𝕜] E} (hT : is_self_adjoint T) :
@@ -254,6 +274,58 @@ begin
254274
exact ext_inner_right 𝕜 (λ y, by simp only [adjoint_inner_left, h x y])
255275
end
256276

277+
@[simp] lemma is_self_adjoint_iff_adjoint_eq (A : E →L[𝕜] E) :
278+
is_self_adjoint (A : E →ₗ[𝕜] E) ↔ A† = A :=
279+
by simp_rw [is_self_adjoint, coe_coe, ← eq_adjoint_iff, eq_comm]
280+
281+
lemma _root_.inner_product_space.is_self_adjoint.adjoint_eq {A : E →L[𝕜] E}
282+
(hA : is_self_adjoint (A : E →ₗ[𝕜] E)) : A† = A :=
283+
by rwa is_self_adjoint_iff_adjoint_eq at hA
284+
285+
lemma _root_.inner_product_space.is_self_adjoint.conj_adjoint {T : E →L[𝕜] E}
286+
(hT : is_self_adjoint (T : E →ₗ[𝕜] E)) (S : E →L[𝕜] F) :
287+
is_self_adjoint (S ∘L T ∘L S† : F →ₗ[𝕜] F) :=
288+
begin
289+
intros x y,
290+
rw [coe_coe, comp_apply, comp_apply, ← adjoint_inner_right, ← coe_coe, hT, coe_coe,
291+
adjoint_inner_left],
292+
refl
293+
end
294+
295+
lemma _root_.inner_product_space.is_self_adjoint.adjoint_conj {T : E →L[𝕜] E}
296+
(hT : is_self_adjoint (T : E →ₗ[𝕜] E)) (S : F →L[𝕜] E) :
297+
is_self_adjoint (S† ∘L T ∘L S : F →ₗ[𝕜] F) :=
298+
begin
299+
convert hT.conj_adjoint (S†),
300+
rw adjoint_adjoint
301+
end
302+
303+
lemma _root_.inner_product_space.is_self_adjoint.conj_orthogonal_projection {T : E →L[𝕜] E}
304+
(hT : is_self_adjoint (T : E →ₗ[𝕜] E)) (U : submodule 𝕜 E) [complete_space U] :
305+
is_self_adjoint (U.subtypeL ∘L orthogonal_projection U ∘L T ∘L U.subtypeL ∘L
306+
orthogonal_projection U : E →ₗ[𝕜] E) :=
307+
begin
308+
have := hT.conj_adjoint (U.subtypeL ∘L orthogonal_projection U),
309+
rwa (orthogonal_projection_is_self_adjoint U).adjoint_eq at this
310+
end
311+
312+
lemma _root_.submodule.adjoint_subtypeL (U : submodule 𝕜 E)
313+
[complete_space U] :
314+
(U.subtypeL)† = orthogonal_projection U :=
315+
begin
316+
symmetry,
317+
rw eq_adjoint_iff,
318+
intros x u,
319+
rw [U.coe_inner, inner_orthogonal_projection_left_eq_right,
320+
orthogonal_projection_mem_subspace_eq_self],
321+
refl
322+
end
323+
324+
lemma _root_.submodule.adjoint_orthogonal_projection (U : submodule 𝕜 E)
325+
[complete_space U] :
326+
(orthogonal_projection U : E →L[𝕜] U)† = U.subtypeL :=
327+
by rw [← U.adjoint_subtypeL, adjoint_adjoint]
328+
257329
/-- `E →L[𝕜] E` is a star algebra with the adjoint as the star operation. -/
258330
instance : has_star (E →L[𝕜] E) := ⟨adjoint⟩
259331
instance : has_involutive_star (E →L[𝕜] E) := ⟨adjoint_adjoint⟩
Lines changed: 126 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,126 @@
1+
/-
2+
Copyright (c) 2022 Anatole Dedecker. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Anatole Dedecker
5+
-/
6+
import analysis.inner_product_space.l2_space
7+
import analysis.inner_product_space.adjoint
8+
9+
/-!
10+
# Positive operators
11+
12+
In this file we define positive operators in a Hilbert space. We follow Bourbaki's choice
13+
of requiring self adjointness in the definition.
14+
15+
## Main definitions
16+
17+
* `is_positive` : a continuous linear map is positive if it is self adjoint and
18+
`∀ x, 0 ≤ re ⟪T x, x⟫`
19+
20+
## Main statements
21+
22+
* `continuous_linear_map.is_positive.conj_adjoint` : if `T : E →L[𝕜] E` is positive,
23+
then for any `S : E →L[𝕜] F`, `S ∘L T ∘L S†` is also positive.
24+
* `continuous_linear_map.is_positive_iff_complex` : in a ***complex*** hilbert space,
25+
checking that `⟪T x, x⟫` is a nonnegative real number for all `x` suffices to prove that
26+
`T` is positive
27+
28+
## References
29+
30+
* [Bourbaki, *Topological Vector Spaces*][bourbaki1987]
31+
32+
## Tags
33+
34+
Positive operator
35+
-/
36+
37+
open inner_product_space is_R_or_C continuous_linear_map
38+
open_locale inner_product complex_conjugate
39+
40+
namespace continuous_linear_map
41+
42+
variables {𝕜 E F : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] [inner_product_space 𝕜 F]
43+
local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y
44+
45+
/-- A continuous linear endomorphism `T` of a Hilbert space is **positive** if it is self adjoint
46+
and `∀ x, 0 ≤ re ⟪T x, x⟫`. -/
47+
def is_positive (T : E →L[𝕜] E) : Prop :=
48+
is_self_adjoint (T : E →ₗ[𝕜] E) ∧ ∀ x, 0 ≤ T.re_apply_inner_self x
49+
50+
lemma is_positive.is_self_adjoint {T : E →L[𝕜] E} (hT : is_positive T) :
51+
is_self_adjoint (T : E →ₗ[𝕜] E) :=
52+
hT.1
53+
54+
lemma is_positive.inner_nonneg_left {T : E →L[𝕜] E} (hT : is_positive T) (x : E) :
55+
0 ≤ re ⟪T x, x⟫ :=
56+
hT.2 x
57+
58+
lemma is_positive.inner_nonneg_right {T : E →L[𝕜] E} (hT : is_positive T) (x : E) :
59+
0 ≤ re ⟪x, T x⟫ :=
60+
by rw inner_re_symm; exact hT.inner_nonneg_left x
61+
62+
lemma is_positive_zero : is_positive (0 : E →L[𝕜] E) :=
63+
begin
64+
refine ⟨is_self_adjoint_zero, λ x, _⟩,
65+
change 0 ≤ re ⟪_, _⟫,
66+
rw [zero_apply, inner_zero_left, zero_hom_class.map_zero]
67+
end
68+
69+
lemma is_positive_id : is_positive (1 : E →L[𝕜] E) :=
70+
⟨λ x y, rfl, λ x, inner_self_nonneg⟩
71+
72+
lemma is_positive.add {T S : E →L[𝕜] E} (hT : T.is_positive)
73+
(hS : S.is_positive) : (T + S).is_positive :=
74+
begin
75+
refine ⟨hT.is_self_adjoint.add hS.is_self_adjoint, λ x, _⟩,
76+
rw [re_apply_inner_self, add_apply, inner_add_left, map_add],
77+
exact add_nonneg (hT.inner_nonneg_left x) (hS.inner_nonneg_left x)
78+
end
79+
80+
lemma is_positive.conj_adjoint [complete_space E] [complete_space F] {T : E →L[𝕜] E}
81+
(hT : T.is_positive) (S : E →L[𝕜] F) : (S ∘L T ∘L S†).is_positive :=
82+
begin
83+
refine ⟨hT.is_self_adjoint.conj_adjoint S, λ x, _⟩,
84+
rw [re_apply_inner_self, comp_apply, ← adjoint_inner_right],
85+
exact hT.inner_nonneg_left _
86+
end
87+
88+
lemma is_positive.adjoint_conj [complete_space E] [complete_space F] {T : E →L[𝕜] E}
89+
(hT : T.is_positive) (S : F →L[𝕜] E) : (S† ∘L T ∘L S).is_positive :=
90+
begin
91+
convert hT.conj_adjoint (S†),
92+
rw adjoint_adjoint
93+
end
94+
95+
lemma is_positive.conj_orthogonal_projection [complete_space E] (U : submodule 𝕜 E) {T : E →L[𝕜] E}
96+
(hT : T.is_positive) [complete_space U] :
97+
(U.subtypeL ∘L orthogonal_projection U ∘L T ∘L U.subtypeL ∘L
98+
orthogonal_projection U).is_positive :=
99+
begin
100+
have := hT.conj_adjoint (U.subtypeL ∘L orthogonal_projection U),
101+
rwa (orthogonal_projection_is_self_adjoint U).adjoint_eq at this
102+
end
103+
104+
lemma is_positive.orthogonal_projection_comp [complete_space E] {T : E →L[𝕜] E}
105+
(hT : T.is_positive) (U : submodule 𝕜 E) [complete_space U] :
106+
(orthogonal_projection U ∘L T ∘L U.subtypeL).is_positive :=
107+
begin
108+
have := hT.conj_adjoint (orthogonal_projection U : E →L[𝕜] U),
109+
rwa [U.adjoint_orthogonal_projection] at this,
110+
end
111+
112+
section complex
113+
114+
variables {E' : Type*} [inner_product_space ℂ E']
115+
116+
lemma is_positive_iff_complex (T : E' →L[ℂ] E') :
117+
is_positive T ↔ ∀ x, (re ⟪T x, x⟫_ℂ : ℂ) = ⟪T x, x⟫_ℂ ∧ 0 ≤ re ⟪T x, x⟫_ℂ :=
118+
begin
119+
simp_rw [is_positive, forall_and_distrib, is_self_adjoint_iff_inner_map_self_real,
120+
eq_conj_iff_re],
121+
refl
122+
end
123+
124+
end complex
125+
126+
end continuous_linear_map

0 commit comments

Comments
 (0)