Documentation

Mathlib.Algebra.Star.Basic

Star monoids, rings, and modules #

We introduce the basic algebraic notions of star monoids, star rings, and star modules. A star algebra is simply a star ring that is also a star module.

These are implemented as "mixin" typeclasses, so to summon a star ring (for example) one needs to write (R : Type*) [Ring R] [StarRing R]. This avoids difficulties with diamond inheritance.

For now we simply do not introduce notations, as different users are expected to feel strongly about the relative merits of r^*, r†, rᘁ, and so on.

Our star rings are actually star non-unital, non-associative, semirings, but of course we can prove star_neg : star (-r) = - star r when the underlying semiring is a ring.

class Star (R : Type u) :

Notation typeclass (with no default notation!) for an algebraic structure with a star operation.

  • star : RR

    A star operation (e.g. complex conjugate).

Instances
class StarMemClass (S : Type u_1) (R : Type u_2) [Star R] [SetLike S R] :

StarMemClass S G states S is a type of subsets s ⊆ G closed under star.

  • star_mem : ∀ {s : S} {r : R}, r sstar r s

    Closure under star.

Instances
theorem StarMemClass.star_mem {S : Type u_1} {R : Type u_2} :
∀ {inst : Star R} {inst_1 : SetLike S R} [self : StarMemClass S R] {s : S} {r : R}, r sstar r s

Closure under star.

instance StarMemClass.instStar {R : Type u} {S : Type w} [Star R] [SetLike S R] [hS : StarMemClass S R] (s : S) :
Star s
Equations
@[simp]
theorem StarMemClass.coe_star {R : Type u} {S : Type w} [Star R] [SetLike S R] [hS : StarMemClass S R] (s : S) (x : s) :
(star x) = star x

Involutive condition.

@[simp]
theorem star_star {R : Type u} [InvolutiveStar R] (r : R) :
star (star r) = r
@[simp]
theorem star_inj {R : Type u} [InvolutiveStar R] {x : R} {y : R} :
star x = star y x = y

star as an equivalence when it is involutive.

Equations
theorem eq_star_of_eq_star {R : Type u} [InvolutiveStar R] {r : R} {s : R} (h : r = star s) :
s = star r
theorem eq_star_iff_eq_star {R : Type u} [InvolutiveStar R] {r : R} {s : R} :
r = star s s = star r
theorem star_eq_iff_star_eq {R : Type u} [InvolutiveStar R] {r : R} {s : R} :
star r = s star s = r
@[simp]
theorem TrivialStar.star_trivial {R : Type u} :
∀ {inst : Star R} [self : TrivialStar R] (r : R), star r = r

Condition that star is trivial

@[simp]
theorem StarMul.star_mul {R : Type u} :
∀ {inst : Mul R} [self : StarMul R] (r s : R), star (r * s) = star s * star r

star skew-distributes over multiplication.

theorem star_star_mul {R : Type u} [Mul R] [StarMul R] (x : R) (y : R) :
star (star x * y) = star y * x
theorem star_mul_star {R : Type u} [Mul R] [StarMul R] (x : R) (y : R) :
star (x * star y) = y * star x
@[simp]
theorem semiconjBy_star_star_star {R : Type u} [Mul R] [StarMul R] {x : R} {y : R} {z : R} :
theorem SemiconjBy.star_star_star {R : Type u} [Mul R] [StarMul R] {x : R} {y : R} {z : R} :
SemiconjBy x y zSemiconjBy (star x) (star z) (star y)

Alias of the reverse direction of semiconjBy_star_star_star.

@[simp]
theorem commute_star_star {R : Type u} [Mul R] [StarMul R] {x : R} {y : R} :
theorem Commute.star_star {R : Type u} [Mul R] [StarMul R] {x : R} {y : R} :
Commute x yCommute (star x) (star y)

Alias of the reverse direction of commute_star_star.

theorem commute_star_comm {R : Type u} [Mul R] [StarMul R] {x : R} {y : R} :
@[simp]
theorem star_mul' {R : Type u} [CommSemigroup R] [StarMul R] (x : R) (y : R) :
star (x * y) = star x * star y

In a commutative ring, make simp prefer leaving the order unchanged.

def starMulEquiv {R : Type u} [Mul R] [StarMul R] :

star as a MulEquiv from R to Rᵐᵒᵖ

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem starMulEquiv_apply {R : Type u} [Mul R] [StarMul R] (x : R) :
starMulEquiv x = MulOpposite.op (star x)
def starMulAut {R : Type u} [CommSemigroup R] [StarMul R] :

star as a MulAut for commutative R.

Equations
  • starMulAut = { toFun := star, invFun := (Function.Involutive.toPerm star ).invFun, left_inv := , right_inv := , map_mul' := }
@[simp]
theorem starMulAut_apply {R : Type u} [CommSemigroup R] [StarMul R] :
∀ (a : R), starMulAut a = star a
@[simp]
theorem star_one (R : Type u) [MulOneClass R] [StarMul R] :
star 1 = 1
@[simp]
theorem star_pow {R : Type u} [Monoid R] [StarMul R] (x : R) (n : ) :
star (x ^ n) = star x ^ n
@[simp]
theorem star_inv {R : Type u} [Group R] [StarMul R] (x : R) :
@[simp]
theorem star_zpow {R : Type u} [Group R] [StarMul R] (x : R) (z : ) :
star (x ^ z) = star x ^ z
@[simp]
theorem star_div {R : Type u} [CommGroup R] [StarMul R] (x : R) (y : R) :
star (x / y) = star x / star y

When multiplication is commutative, star preserves division.

@[reducible, inline]
abbrev starMulOfComm {R : Type u_1} [CommMonoid R] :

Any commutative monoid admits the trivial *-structure.

See note [reducible non-instances].

Equations
theorem star_id_of_comm {R : Type u_1} [CommSemiring R] {x : R} :
star x = x

Note that since starMulOfComm is reducible, simp can already prove this.

@[simp]
theorem StarAddMonoid.star_add {R : Type u} :
∀ {inst : AddMonoid R} [self : StarAddMonoid R] (r s : R), star (r + s) = star r + star s

star commutes with addition

def starAddEquiv {R : Type u} [AddMonoid R] [StarAddMonoid R] :
R ≃+ R

star as an AddEquiv

Equations
  • starAddEquiv = { toFun := star, invFun := (Function.Involutive.toPerm star ).invFun, left_inv := , right_inv := , map_add' := }
@[simp]
theorem starAddEquiv_apply {R : Type u} [AddMonoid R] [StarAddMonoid R] :
∀ (a : R), starAddEquiv a = star a
@[simp]
theorem star_zero (R : Type u) [AddMonoid R] [StarAddMonoid R] :
star 0 = 0
@[simp]
theorem star_eq_zero {R : Type u} [AddMonoid R] [StarAddMonoid R] {x : R} :
star x = 0 x = 0
theorem star_ne_zero {R : Type u} [AddMonoid R] [StarAddMonoid R] {x : R} :
star x 0 x 0
@[simp]
theorem star_neg {R : Type u} [AddGroup R] [StarAddMonoid R] (r : R) :
star (-r) = -star r
@[simp]
theorem star_sub {R : Type u} [AddGroup R] [StarAddMonoid R] (r : R) (s : R) :
star (r - s) = star r - star s
@[simp]
theorem star_nsmul {R : Type u} [AddMonoid R] [StarAddMonoid R] (n : ) (x : R) :
star (n x) = n star x
@[simp]
theorem star_zsmul {R : Type u} [AddGroup R] [StarAddMonoid R] (n : ) (x : R) :
star (n x) = n star x
theorem StarRing.star_add {R : Type u} :
∀ {inst : NonUnitalNonAssocSemiring R} [self : StarRing R] (r s : R), star (r + s) = star r + star s

star commutes with addition

@[instance 100]
Equations

star as a RingEquiv from R to Rᵐᵒᵖ

Equations
  • starRingEquiv = { toFun := fun (x : R) => MulOpposite.op (star x), invFun := (starAddEquiv.trans MulOpposite.opAddEquiv).invFun, left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
theorem starRingEquiv_apply {R : Type u} [NonUnitalNonAssocSemiring R] [StarRing R] (x : R) :
starRingEquiv x = MulOpposite.op (star x)
@[simp]
theorem star_natCast {R : Type u} [NonAssocSemiring R] [StarRing R] (n : ) :
star n = n
@[simp]
theorem star_ofNat {R : Type u} [NonAssocSemiring R] [StarRing R] (n : ) [n.AtLeastTwo] :
@[simp]
theorem star_intCast {R : Type u} [Ring R] [StarRing R] (z : ) :
star z = z

star as a ring automorphism, for commutative R.

Equations
  • starRingAut = { toFun := star, invFun := starAddEquiv.invFun, left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
theorem starRingAut_apply {R : Type u} [CommSemiring R] [StarRing R] :
∀ (a : R), starRingAut a = star a
def starRingEnd (R : Type u) [CommSemiring R] [StarRing R] :
R →+* R

star as a ring endomorphism, for commutative R. This is used to denote complex conjugation, and is available under the notation conj in the locale ComplexConjugate.

Note that this is the preferred form (over starRingAut, available under the same hypotheses) because the notation E →ₗ⋆[R] F for an R-conjugate-linear map (short for E →ₛₗ[starRingEnd R] F) does not pretty-print if there is a coercion involved, as would be the case for (↑starRingAut : R →* R).

Equations
Instances For

star as a ring endomorphism, for commutative R. This is used to denote complex conjugation, and is available under the notation conj in the locale ComplexConjugate.

Note that this is the preferred form (over starRingAut, available under the same hypotheses) because the notation E →ₗ⋆[R] F for an R-conjugate-linear map (short for E →ₛₗ[starRingEnd R] F) does not pretty-print if there is a coercion involved, as would be the case for (↑starRingAut : R →* R).

Equations
theorem starRingEnd_apply {R : Type u} [CommSemiring R] [StarRing R] (x : R) :

This is not a simp lemma, since we usually want simp to keep starRingEnd bundled. For example, for complex conjugation, we don't want simp to turn conj x into the bare function star x automatically since most lemmas are about conj x.

theorem starRingEnd_self_apply {R : Type u} [CommSemiring R] [StarRing R] (x : R) :
(starRingEnd R) ((starRingEnd R) x) = x
Equations
theorem RingHom.star_def {R : Type u} [CommSemiring R] [StarRing R] {S : Type u_1} [NonAssocSemiring S] (f : S →+* R) :
star f = (starRingEnd R).comp f
theorem RingHom.star_apply {R : Type u} [CommSemiring R] [StarRing R] {S : Type u_1} [NonAssocSemiring S] (f : S →+* R) (s : S) :
(star f) s = star (f s)
theorem Complex.conj_conj {R : Type u} [CommSemiring R] [StarRing R] (x : R) :
(starRingEnd R) ((starRingEnd R) x) = x

Alias of starRingEnd_self_apply.

theorem RCLike.conj_conj {R : Type u} [CommSemiring R] [StarRing R] (x : R) :
(starRingEnd R) ((starRingEnd R) x) = x

Alias of starRingEnd_self_apply.

@[simp]
theorem conj_trivial {R : Type u} [CommSemiring R] [StarRing R] [TrivialStar R] (a : R) :
(starRingEnd R) a = a
@[simp]
theorem star_inv' {R : Type u} [DivisionSemiring R] [StarRing R] (x : R) :
@[simp]
theorem star_zpow₀ {R : Type u} [DivisionSemiring R] [StarRing R] (x : R) (z : ) :
star (x ^ z) = star x ^ z
@[simp]
theorem star_div' {R : Type u} [Semifield R] [StarRing R] (x : R) (y : R) :
star (x / y) = star x / star y

When multiplication is commutative, star preserves division.

@[reducible, inline]
abbrev starRingOfComm {R : Type u_1} [CommSemiring R] :

Any commutative semiring admits the trivial *-structure.

See note [reducible non-instances].

Equations
Equations
Equations
class StarModule (R : Type u) (A : Type v) [Star R] [Star A] [SMul R A] :

A star module A over a star ring R is a module which is a star add monoid, and the two star structures are compatible in the sense star (r • a) = star r • star a.

Note that it is up to the user of this typeclass to enforce [Semiring R] [StarRing R] [AddCommMonoid A] [StarAddMonoid A] [Module R A], and that the statement only requires [Star R] [Star A] [SMul R A].

If used as [CommRing R] [StarRing R] [Semiring A] [StarRing A] [Algebra R A], this represents a star algebra.

Instances
@[simp]
theorem StarModule.star_smul {R : Type u} {A : Type v} :
∀ {inst : Star R} {inst_1 : Star A} {inst_2 : SMul R A} [self : StarModule R A] (r : R) (a : A), star (r a) = star r star a

star commutes with scalar multiplication

instance StarMul.toStarModule {R : Type u} [CommMonoid R] [StarMul R] :

A commutative star monoid is a star module over itself via Monoid.toMulAction.

Equations
  • =
Equations
  • =
Equations
  • =

Instance needed to define star-linear maps over a commutative star ring (ex: conjugate-linear maps when R = ℂ).

Equations
  • =
class StarHomClass (F : Type u_1) (R : outParam (Type u_2)) (S : outParam (Type u_3)) [Star R] [Star S] [FunLike F R S] :

StarHomClass F R S states that F is a type of star-preserving maps from R to S.

  • map_star : ∀ (f : F) (r : R), f (star r) = star (f r)

    the maps preserve star

Instances
theorem StarHomClass.map_star {F : Type u_1} {R : outParam (Type u_2)} {S : outParam (Type u_3)} :
∀ {inst : Star R} {inst_1 : Star S} {inst_2 : FunLike F R S} [self : StarHomClass F R S] (f : F) (r : R), f (star r) = star (f r)

the maps preserve star

Instances #

instance Units.instStarMul {R : Type u} [Monoid R] [StarMul R] :
Equations
@[simp]
theorem Units.coe_star {R : Type u} [Monoid R] [StarMul R] (u : Rˣ) :
(star u) = star u
@[simp]
theorem Units.coe_star_inv {R : Type u} [Monoid R] [StarMul R] (u : Rˣ) :
(star u)⁻¹ = star u⁻¹
instance Units.instStarModule {R : Type u} [Monoid R] [StarMul R] {A : Type u_1} [Star A] [SMul R A] [StarModule R A] :
Equations
  • =
theorem IsUnit.star {R : Type u} [Monoid R] [StarMul R] {a : R} :
IsUnit aIsUnit (star a)
@[simp]
theorem isUnit_star {R : Type u} [Monoid R] [StarMul R] {a : R} :
instance Invertible.star {R : Type u_1} [MulOneClass R] [StarMul R] (r : R) [Invertible r] :
Equations
theorem star_invOf {R : Type u_1} [Monoid R] [StarMul R] (r : R) [Invertible r] [Invertible (star r)] :
theorem IsLeftRegular.star {R : Type u} [Mul R] [StarMul R] {x : R} (hx : IsLeftRegular x) :
theorem IsRightRegular.star {R : Type u} [Mul R] [StarMul R] {x : R} (hx : IsRightRegular x) :
theorem IsRegular.star {R : Type u} [Mul R] [StarMul R] {x : R} (hx : IsRegular x) :
@[simp]
@[simp]
@[simp]
theorem isRegular_star_iff {R : Type u} [Mul R] [StarMul R] {x : R} :

The opposite type carries the same star operation.

Equations
@[simp]
theorem MulOpposite.op_star {R : Type u} [Star R] (r : R) :
Equations
Equations
Equations
Equations

A commutative star monoid is a star module over its opposite via Monoid.toOppositeMulAction.

Equations
  • =