Self-adjoint, skew-adjoint and normal elements of a star additive group #
This file defines selfAdjoint R (resp. skewAdjoint R), where R is a star additive group,
as the additive subgroup containing the elements that satisfy star x = x (resp. star x = -x).
This includes, for instance, (skew-)Hermitian operators on Hilbert spaces.
We also define IsStarNormal R, a Prop that states that an element x satisfies
star x * x = x * star x.
Implementation notes #
- When Ris aStarModule R₂ R, thenselfAdjoint Rhas a naturalModule (selfAdjoint R₂) (selfAdjoint R)structure. However, doing this literally would be undesirable since in the main case of interest (R₂ = ℂ) we wantModule ℝ (selfAdjoint R)and notModule (selfAdjoint ℂ) (selfAdjoint R). We solve this issue by adding the typeclass[TrivialStar R₃], of whichℝis an instance (registered inData/Real/Basic), and then add a[Module R₃ (selfAdjoint R)]instance whenever we have[Module R₃ R] [TrivialStar R₃]. (Another approach would have been to define[StarInvariantScalars R₃ R]to express the fact thatstar (x • v) = x • star v, but this typeclass would have the disadvantage of taking two type arguments.)
TODO #
- Define IsSkewAdjointto matchIsSelfAdjoint.
- Define fun z x => z * x * star z(i.e. conjugation byz) as a monoid action ofRonR(similar to the existingConjActfor groups), and then state the fact thatselfAdjoint Ris invariant under it.
An element is self-adjoint if it is equal to its star.
Equations
- IsSelfAdjoint x = (star x = x)
Instances For
A normal element of a star monoid commutes with its adjoint.
All elements are self-adjoint when star is trivial.
Self-adjoint elements commute if and only if their product is self-adjoint.
Functions in a StarHomClass preserve self-adjoint elements.
The self-adjoint elements of a star additive group, as an additive subgroup.
Equations
- selfAdjoint R = { carrier := {x : R | IsSelfAdjoint x}, add_mem' := ⋯, zero_mem' := ⋯, neg_mem' := ⋯ }
Instances For
The skew-adjoint elements of a star additive group, as an additive subgroup.
Equations
- skewAdjoint R = { carrier := {x : R | star x = -x}, add_mem' := ⋯, zero_mem' := ⋯, neg_mem' := ⋯ }
Instances For
Equations
- selfAdjoint.instInhabitedSubtypeMemAddSubgroup = { default := 0 }
Equations
- ⋯ = ⋯
Equations
- selfAdjoint.instOneSubtypeMemAddSubgroup = { one := ⟨1, ⋯⟩ }
Equations
- ⋯ = ⋯
Equations
- selfAdjoint.instPowSubtypeMemAddSubgroupNat = { pow := fun (x : ↥(selfAdjoint R)) (n : ℕ) => ⟨↑x ^ n, ⋯⟩ }
Equations
- selfAdjoint.instMulSubtypeMemAddSubgroup = { mul := fun (x y : ↥(selfAdjoint R)) => ⟨↑x * ↑y, ⋯⟩ }
Equations
- selfAdjoint.instCommRingSubtypeMemAddSubgroup = Function.Injective.commRing (fun (a : ↥(selfAdjoint R)) => ↑a) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- selfAdjoint.instInvSubtypeMemAddSubgroup = { inv := fun (x : ↥(selfAdjoint R)) => ⟨(↑x)⁻¹, ⋯⟩ }
Equations
- selfAdjoint.instDivSubtypeMemAddSubgroup = { div := fun (x y : ↥(selfAdjoint R)) => ⟨↑x / ↑y, ⋯⟩ }
Equations
- selfAdjoint.instPowSubtypeMemAddSubgroupInt = { pow := fun (x : ↥(selfAdjoint R)) (z : ℤ) => ⟨↑x ^ z, ⋯⟩ }
Equations
- selfAdjoint.instSMulNNRat = { smul := fun (a : ℚ≥0) (x : ↥(selfAdjoint R)) => ⟨a • ↑x, ⋯⟩ }
Equations
- selfAdjoint.instSMulRat = { smul := fun (a : ℚ) (x : ↥(selfAdjoint R)) => ⟨a • ↑x, ⋯⟩ }
Equations
- selfAdjoint.instField = Function.Injective.field (fun (a : ↥(selfAdjoint R)) => ↑a) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- selfAdjoint.instSMulSubtypeMemAddSubgroupOfStarModule = { smul := fun (r : R) (x : ↥(selfAdjoint A)) => ⟨r • ↑x, ⋯⟩ }
Equations
- selfAdjoint.instMulActionSubtypeMemAddSubgroupOfStarModule = Function.Injective.mulAction Subtype.val ⋯ ⋯
Equations
- selfAdjoint.instDistribMulActionSubtypeMemAddSubgroupOfStarModule = Function.Injective.distribMulAction (selfAdjoint A).subtype ⋯ ⋯
Equations
- selfAdjoint.instModuleSubtypeMemAddSubgroupOfStarModule = Function.Injective.module R (selfAdjoint A).subtype ⋯ ⋯
Equations
- skewAdjoint.instInhabitedSubtypeMemAddSubgroup = { default := 0 }
Equations
- ⋯ = ⋯
Equations
- skewAdjoint.instSMulSubtypeMemAddSubgroupOfStarModule = { smul := fun (r : R) (x : ↥(skewAdjoint A)) => ⟨r • ↑x, ⋯⟩ }
Equations
- skewAdjoint.instDistribMulActionSubtypeMemAddSubgroupOfStarModule = Function.Injective.distribMulAction (skewAdjoint A).subtype ⋯ ⋯
Equations
- skewAdjoint.instModuleSubtypeMemAddSubgroupOfStarModule = Function.Injective.module R (skewAdjoint A).subtype ⋯ ⋯
Scalar multiplication of a self-adjoint element by a skew-adjoint element produces a skew-adjoint element.
Scalar multiplication of a skew-adjoint element by a skew-adjoint element produces a self-adjoint element.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Alias of the forward direction of Pi.isSelfAdjoint.