Skip to content

Commit

Permalink
feat(analysis/inner_product_space/adjoint): `is_self_adjoint_iff_eq_a… (
Browse files Browse the repository at this point in the history
#12047)

…djoint`

A self-adjoint linear map is equal to its adjoint.
  • Loading branch information
hparshall committed Feb 15, 2022
1 parent 92ac8ff commit a1283d0
Showing 1 changed file with 4 additions and 0 deletions.
4 changes: 4 additions & 0 deletions src/analysis/inner_product_space/adjoint.lean
Expand Up @@ -273,6 +273,10 @@ begin
refine ext_inner_right_basis b (λ i, by simp only [h i, adjoint_inner_left]),
end

lemma is_self_adjoint_iff_eq_adjoint (A : E →ₗ[𝕜] E) :
is_self_adjoint A ↔ A = A.adjoint :=
by rw [is_self_adjoint, ← linear_map.eq_adjoint_iff]

/-- `E →ₗ[𝕜] E` is a star algebra with the adjoint as the star operation. -/
instance : has_star (E →ₗ[𝕜] E) := ⟨adjoint⟩
instance : has_involutive_star (E →ₗ[𝕜] E) := ⟨adjoint_adjoint⟩
Expand Down

0 comments on commit a1283d0

Please sign in to comment.