Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - refactor(Algebra/Star/*): Allow for star operation on non-associative algebras #6562

Closed
wants to merge 21 commits into from
Closed
Show file tree
Hide file tree
Changes from 18 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
5 changes: 3 additions & 2 deletions Mathlib/Algebra/Algebra/Unitization.lean
Original file line number Diff line number Diff line change
Expand Up @@ -578,8 +578,9 @@ instance instStarModule [CommSemiring R] [StarRing R] [AddCommMonoid A] [StarAdd
[Module R A] [StarModule R A] : StarModule R (Unitization R A) where
star_smul r x := ext (by simp) (by simp)

instance instStarRing [CommSemiring R] [StarRing R] [NonUnitalSemiring A] [StarRing A] [Module R A]
[IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] : StarRing (Unitization R A) :=
instance instStarRing [CommSemiring R] [StarRing R] [NonUnitalNonAssocSemiring A] [StarRing A]
[Module R A] [StarModule R A] :
StarRing (Unitization R A) :=
{ Unitization.instStarAddMonoid with
star_mul := fun x y =>
ext (by simp [-star_mul']) (by simp [-star_mul', add_comm (star x.fst • star y.snd)]) }
Expand Down
76 changes: 38 additions & 38 deletions Mathlib/Algebra/Star/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,7 +25,7 @@ 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 semirings, but of course we can prove
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.
-/

Expand Down Expand Up @@ -121,21 +121,21 @@ export TrivialStar (star_trivial)

attribute [simp] star_trivial

/-- A `*`-semigroup is a semigroup `R` with an involutive operation `star`
such that `star (r * s) = star s * star r`.
/-- `StarMul R`, for a type `R` with a multiplication, provides `R` with an involutive operation
`star` such that `star (r * s) = star s * star r`.
-/
class StarSemigroup (R : Type u) [Semigroup R] extends InvolutiveStar R where
class StarMul (R : Type u) [Mul R] extends InvolutiveStar R where
/-- `star` skew-distributes over multiplication. -/
star_mul : ∀ r s : R, star (r * s) = star s * star r
#align star_semigroup StarSemigroup
#align star_semigroup StarMul

export StarSemigroup (star_mul)
export StarMul (star_mul)

attribute [simp 900] star_mul

section StarSemigroup
section StarMul

variable [Semigroup R] [StarSemigroup R]
variable [Mul R] [StarMul R]

theorem star_star_mul (x y : R) : star (star x * y) = star y * x := by rw [star_mul, star_star]
#align star_star_mul star_star_mul
Expand Down Expand Up @@ -164,17 +164,17 @@ theorem commute_star_comm {x y : R} : Commute (star x) y ↔ Commute x (star y)
rw [← commute_star_star, star_star]
#align commute_star_comm commute_star_comm

end StarSemigroup
end StarMul

/-- In a commutative ring, make `simp` prefer leaving the order unchanged. -/
@[simp]
theorem star_mul' [CommSemigroup R] [StarSemigroup R] (x y : R) : star (x * y) = star x * star y :=
theorem star_mul' [CommSemigroup R] [StarMul R] (x y : R) : star (x * y) = star x * star y :=
(star_mul x y).trans (mul_comm _ _)
#align star_mul' star_mul'

/-- `star` as a `MulEquiv` from `R` to `Rᵐᵒᵖ` -/
@[simps apply]
def starMulEquiv [Semigroup R] [StarSemigroup R] : R ≃* Rᵐᵒᵖ :=
def starMulEquiv [Mul R] [StarMul R] : R ≃* Rᵐᵒᵖ :=
{ (InvolutiveStar.star_involutive.toPerm star).trans opEquiv with
toFun := fun x => MulOpposite.op (star x)
map_mul' := fun x y => by simp only [star_mul, op_mul] }
Expand All @@ -183,7 +183,7 @@ def starMulEquiv [Semigroup R] [StarSemigroup R] : R ≃* Rᵐᵒᵖ :=

/-- `star` as a `MulAut` for commutative `R`. -/
@[simps apply]
def starMulAut [CommSemigroup R] [StarSemigroup R] : MulAut R :=
def starMulAut [CommSemigroup R] [StarMul R] : MulAut R :=
{ InvolutiveStar.star_involutive.toPerm star with
toFun := star
map_mul' := star_mul' }
Expand All @@ -193,32 +193,32 @@ def starMulAut [CommSemigroup R] [StarSemigroup R] : MulAut R :=
variable (R)

@[simp]
theorem star_one [Monoid R] [StarSemigroup R] : star (1 : R) = 1 :=
theorem star_one [MulOneClass R] [StarMul R] : star (1 : R) = 1 :=
op_injective <| (starMulEquiv : R ≃* Rᵐᵒᵖ).map_one.trans (op_one _).symm
#align star_one star_one

variable {R}

@[simp]
theorem star_pow [Monoid R] [StarSemigroup R] (x : R) (n : ℕ) : star (x ^ n) = star x ^ n :=
theorem star_pow [Monoid R] [StarMul R] (x : R) (n : ℕ) : star (x ^ n) = star x ^ n :=
op_injective <|
((starMulEquiv : R ≃* Rᵐᵒᵖ).toMonoidHom.map_pow x n).trans (op_pow (star x) n).symm
#align star_pow star_pow

@[simp]
theorem star_inv [Group R] [StarSemigroup R] (x : R) : star x⁻¹ = (star x)⁻¹ :=
theorem star_inv [Group R] [StarMul R] (x : R) : star x⁻¹ = (star x)⁻¹ :=
op_injective <| ((starMulEquiv : R ≃* Rᵐᵒᵖ).toMonoidHom.map_inv x).trans (op_inv (star x)).symm
#align star_inv star_inv

@[simp]
theorem star_zpow [Group R] [StarSemigroup R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z :=
theorem star_zpow [Group R] [StarMul R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z :=
op_injective <|
((starMulEquiv : R ≃* Rᵐᵒᵖ).toMonoidHom.map_zpow x z).trans (op_zpow (star x) z).symm
#align star_zpow star_zpow

/-- When multiplication is commutative, `star` preserves division. -/
@[simp]
theorem star_div [CommGroup R] [StarSemigroup R] (x y : R) : star (x / y) = star x / star y :=
theorem star_div [CommGroup R] [StarMul R] (x y : R) : star (x / y) = star x / star y :=
map_div (starMulAut : R ≃* R) _ _
#align star_div star_div

Expand All @@ -227,17 +227,17 @@ theorem star_div [CommGroup R] [StarSemigroup R] (x y : R) : star (x / y) = star
See note [reducible non-instances].
-/
@[reducible]
def starSemigroupOfComm {R : Type*} [CommMonoid R] : StarSemigroup R where
def starMulOfComm {R : Type*} [CommMonoid R] : StarMul R where
star := id
star_involutive _ := rfl
star_mul := mul_comm
#align star_semigroup_of_comm starSemigroupOfComm
#align star_semigroup_of_comm starMulOfComm

section

attribute [local instance] starSemigroupOfComm
attribute [local instance] starMulOfComm

/-- Note that since `starSemigroupOfComm` is reducible, `simp` can already prove this. -/
/-- Note that since `starMulOfComm` is reducible, `simp` can already prove this. -/
theorem star_id_of_comm {R : Type*} [CommSemiring R] {x : R} : star x = x :=
rfl
#align star_id_of_comm star_id_of_comm
Expand Down Expand Up @@ -302,35 +302,35 @@ theorem star_zsmul [AddGroup R] [StarAddMonoid R] (x : R) (n : ℤ) : star (n
(starAddEquiv : R ≃+ R).toAddMonoidHom.map_zsmul _ _
#align star_zsmul star_zsmul

/-- A `*`-ring `R` is a (semi)ring with an involutive `star` operation which is additive
which makes `R` with its multiplicative structure into a `*`-semigroup
/-- A `*`-ring `R` is a non-unital, non-associative (semi)ring with an involutive `star` operation
which is additive which makes `R` with its multiplicative structure into a `*`-multiplication
(i.e. `star (r * s) = star s * star r`). -/
class StarRing (R : Type u) [NonUnitalSemiring R] extends StarSemigroup R where
class StarRing (R : Type u) [NonUnitalNonAssocSemiring R] extends StarMul R where
/-- `star` commutes with addition -/
star_add : ∀ r s : R, star (r + s) = star r + star s
#align star_ring StarRing

instance (priority := 100) StarRing.toStarAddMonoid [NonUnitalSemiring R] [StarRing R] :
instance (priority := 100) StarRing.toStarAddMonoid [NonUnitalNonAssocSemiring R] [StarRing R] :
StarAddMonoid R where
star_add := StarRing.star_add
#align star_ring.to_star_add_monoid StarRing.toStarAddMonoid

/-- `star` as a `RingEquiv` from `R` to `Rᵐᵒᵖ` -/
@[simps apply]
def starRingEquiv [NonUnitalSemiring R] [StarRing R] : R ≃+* Rᵐᵒᵖ :=
def starRingEquiv [NonUnitalNonAssocSemiring R] [StarRing R] : R ≃+* Rᵐᵒᵖ :=
{ starAddEquiv.trans (MulOpposite.opAddEquiv : R ≃+ Rᵐᵒᵖ), starMulEquiv with
toFun := fun x => MulOpposite.op (star x) }
#align star_ring_equiv starRingEquiv
#align star_ring_equiv_apply starRingEquiv_apply

@[simp, norm_cast]
theorem star_natCast [Semiring R] [StarRing R] (n : ℕ) : star (n : R) = n :=
theorem star_natCast [NonAssocSemiring R] [StarRing R] (n : ℕ) : star (n : R) = n :=
(congr_arg unop (map_natCast (starRingEquiv : R ≃+* Rᵐᵒᵖ) n)).trans (unop_natCast _)
#align star_nat_cast star_natCast

--Porting note: new theorem
@[simp]
theorem star_ofNat [Semiring R] [StarRing R] (n : ℕ) [n.AtLeastTwo] :
theorem star_ofNat [NonAssocSemiring R] [StarRing R] (n : ℕ) [n.AtLeastTwo] :
star (no_index (OfNat.ofNat n) : R) = OfNat.ofNat n :=
star_natCast _

Expand Down Expand Up @@ -459,7 +459,7 @@ See note [reducible non-instances].
-/
@[reducible]
def starRingOfComm {R : Type*} [CommSemiring R] : StarRing R :=
{ starSemigroupOfComm with
{ starMulOfComm with
star := id
star_add := fun _ _ => rfl }
#align star_ring_of_comm starRingOfComm
Expand All @@ -486,7 +486,7 @@ export StarModule (star_smul)
attribute [simp] star_smul

/-- A commutative star monoid is a star module over itself via `Monoid.toMulAction`. -/
instance StarSemigroup.to_starModule [CommMonoid R] [StarSemigroup R] : StarModule R R :=
instance StarSemigroup.to_starModule [CommMonoid R] [StarMul R] : StarModule R R :=
⟨star_mul'⟩
#align star_semigroup.to_star_module StarSemigroup.to_starModule

Expand Down Expand Up @@ -517,9 +517,9 @@ end

namespace Units

variable [Monoid R] [StarSemigroup R]
variable [Monoid R] [StarMul R]

instance : StarSemigroup Rˣ where
instance : StarMul Rˣ where
star u :=
{ val := star u
inv := star ↑u⁻¹
Expand All @@ -543,12 +543,12 @@ instance {A : Type*} [Star A] [SMul R A] [StarModule R A] : StarModule Rˣ A :=

end Units

theorem IsUnit.star [Monoid R] [StarSemigroup R] {a : R} : IsUnit a → IsUnit (star a)
theorem IsUnit.star [Monoid R] [StarMul R] {a : R} : IsUnit a → IsUnit (star a)
| ⟨u, hu⟩ => ⟨Star.star u, hu ▸ rfl⟩
#align is_unit.star IsUnit.star

@[simp]
theorem isUnit_star [Monoid R] [StarSemigroup R] {a : R} : IsUnit (star a) ↔ IsUnit a :=
theorem isUnit_star [Monoid R] [StarMul R] {a : R} : IsUnit (star a) ↔ IsUnit a :=
⟨fun h => star_star a ▸ h.star, IsUnit.star⟩
#align is_unit_star isUnit_star

Expand All @@ -560,14 +560,14 @@ theorem Ring.inverse_star [Semiring R] [StarRing R] (a : R) :
rw [Ring.inverse_non_unit _ ha, Ring.inverse_non_unit _ (mt isUnit_star.mp ha), star_zero]
#align ring.inverse_star Ring.inverse_star

instance Invertible.star {R : Type*} [Monoid R] [StarSemigroup R] (r : R) [Invertible r] :
instance Invertible.star {R : Type*} [MulOneClass R] [StarMul R] (r : R) [Invertible r] :
Invertible (star r) where
invOf := Star.star (⅟ r)
invOf_mul_self := by rw [← star_mul, mul_invOf_self, star_one]
mul_invOf_self := by rw [← star_mul, invOf_mul_self, star_one]
#align invertible.star Invertible.star

theorem star_invOf {R : Type*} [Monoid R] [StarSemigroup R] (r : R) [Invertible r]
theorem star_invOf {R : Type*} [Monoid R] [StarMul R] (r : R) [Invertible r]
[Invertible (star r)] : star (⅟ r) = ⅟ (star r) := by
have : star (⅟ r) = star (⅟ r) * ((star r) * ⅟ (star r)) := by
simp only [mul_invOf_self, mul_one]
Expand All @@ -594,7 +594,7 @@ theorem op_star [Star R] (r : R) : op (star r) = star (op r) :=
instance [InvolutiveStar R] : InvolutiveStar Rᵐᵒᵖ where
star_involutive r := unop_injective (star_star r.unop)

instance [Monoid R] [StarSemigroup R] : StarSemigroup Rᵐᵒᵖ where
instance [Mul R] [StarMul R] : StarMul Rᵐᵒᵖ where
star_mul x y := unop_injective (star_mul y.unop x.unop)

instance [AddMonoid R] [StarAddMonoid R] : StarAddMonoid Rᵐᵒᵖ where
Expand All @@ -607,7 +607,7 @@ end MulOpposite

/-- A commutative star monoid is a star module over its opposite via
`Monoid.toOppositeMulAction`. -/
instance StarSemigroup.toOpposite_starModule [CommMonoid R] [StarSemigroup R] :
instance StarSemigroup.toOpposite_starModule [CommMonoid R] [StarMul R] :
StarModule Rᵐᵒᵖ R :=
⟨fun r s => star_mul' s r.unop⟩
#align star_semigroup.to_opposite_star_module StarSemigroup.toOpposite_starModule
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Star/BigOperators.lean
Original file line number Diff line number Diff line change
Expand Up @@ -19,7 +19,7 @@ variable {R : Type*}
open BigOperators

@[simp]
theorem star_prod [CommMonoid R] [StarSemigroup R] {α : Type*} (s : Finset α) (f : α → R) :
theorem star_prod [CommMonoid R] [StarMul R] {α : Type*} (s : Finset α) (f : α → R) :
star (∏ x in s, f x) = ∏ x in s, star (f x) := map_prod (starMulAut : R ≃* R) _ _
#align star_prod star_prod

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Star/CHSH.lean
Original file line number Diff line number Diff line change
Expand Up @@ -82,7 +82,7 @@ The physical interpretation is that `A₀` and `A₁` are a pair of boolean obse
are spacelike separated from another pair `B₀` and `B₁` of boolean observables.
-/
--@[nolint has_nonempty_instance] Porting note: linter does not exist
structure IsCHSHTuple {R} [Monoid R] [StarSemigroup R] (A₀ A₁ B₀ B₁ : R) : Prop where
structure IsCHSHTuple {R} [Monoid R] [StarMul R] (A₀ A₁ B₀ B₁ : R) : Prop where
A₀_inv : A₀ ^ 2 = 1
A₁_inv : A₁ ^ 2 = 1
B₀_inv : B₀ ^ 2 = 1
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Star/Center.lean
Original file line number Diff line number Diff line change
Expand Up @@ -10,7 +10,7 @@ import Mathlib.Algebra.Star.Pointwise

/-! # `Set.center`, `Set.centralizer` and the `star` operation -/

variable {R : Type*} [Semigroup R] [StarSemigroup R] {a : R} {s : Set R}
variable {R : Type*} [Mul R] [StarMul R] {a : R} {s : Set R}

theorem Set.star_mem_center (ha : a ∈ Set.center R) : star a ∈ Set.center R := by
simpa only [star_mul, star_star] using fun g =>
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Star/Free.lean
Original file line number Diff line number Diff line change
Expand Up @@ -23,7 +23,7 @@ namespace FreeMonoid

variable {α : Type*}

instance : StarSemigroup (FreeMonoid α) where
instance : StarMul (FreeMonoid α) where
star := List.reverse
star_involutive := List.reverse_reverse
star_mul := List.reverse_append
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Algebra/Star/Module.lean
Original file line number Diff line number Diff line change
Expand Up @@ -82,7 +82,7 @@ def starLinearEquiv (R : Type*) {A : Type*} [CommSemiring R] [StarRing R] [AddCo
map_smul' := star_smul }
#align star_linear_equiv starLinearEquiv

variable (R : Type*) (A : Type*) [Semiring R] [StarSemigroup R] [TrivialStar R] [AddCommGroup A]
variable (R : Type*) (A : Type*) [Semiring R] [StarMul R] [TrivialStar R] [AddCommGroup A]
[Module R A] [StarAddMonoid A] [StarModule R A]

/-- The self-adjoint elements of a star module, as a submodule. -/
Expand Down Expand Up @@ -184,7 +184,7 @@ def StarModule.decomposeProdAdjoint : A ≃ₗ[R] selfAdjoint A × skewAdjoint A

@[simp]
theorem algebraMap_star_comm {R A : Type*} [CommSemiring R] [StarRing R] [Semiring A]
[StarSemigroup A] [Algebra R A] [StarModule R A] (r : R) :
[StarMul A] [Algebra R A] [StarModule R A] (r : R) :
algebraMap R A (star r) = star (algebraMap R A r) := by
simp only [Algebra.algebraMap_eq_smul_one, star_smul, star_one]
#align algebra_map_star_comm algebraMap_star_comm
18 changes: 10 additions & 8 deletions Mathlib/Algebra/Star/NonUnitalSubalgebra.lean
Original file line number Diff line number Diff line change
Expand Up @@ -26,10 +26,11 @@ instance instInvolutiveStar {S R : Type*} [InvolutiveStar R] [SetLike S R] [Star
(s : S) : InvolutiveStar s where
star_involutive r := Subtype.ext <| star_star (r : R)

/-- In a star semigroup (i.e., a semigroup with an antimultiplicative involutive star operation),
any star-closed subset which is also closed under multiplication is itself a star semigroup. -/
instance instStarSemigroup {S R : Type*} [Semigroup R] [StarSemigroup R] [SetLike S R]
[MulMemClass S R] [StarMemClass S R] (s : S) : StarSemigroup s where
/-- In a star multiplication (i.e., a multiplication with an antimultiplicative involutive star
operation), any star-closed subset which is also closed under multiplication is itself a star
multiplication. -/
mans0954 marked this conversation as resolved.
Show resolved Hide resolved
instance instStarMul {S R : Type*} [Mul R] [StarMul R] [SetLike S R]
[MulMemClass S R] [StarMemClass S R] (s : S) : StarMul s where
star_mul _ _ := Subtype.ext <| star_mul _ _

/-- In a `StarAddMonoid` (i.e., an additive monoid with an additive involutive star operation), any
Expand All @@ -39,11 +40,12 @@ instance instStarAddMonoid {S R : Type*} [AddMonoid R] [StarAddMonoid R] [SetLik
[AddSubmonoidClass S R] [StarMemClass S R] (s : S) : StarAddMonoid s where
star_add _ _ := Subtype.ext <| star_add _ _

/-- In a star ring (i.e., a non-unital semiring with an additive, antimultiplicative, involutive
star operation), an star-closed non-unital subsemiring is itself a star ring. -/
instance instStarRing {S R : Type*} [NonUnitalSemiring R] [StarRing R] [SetLike S R]
/-- In a star ring (i.e., a non-unital, non-associative, semiring with an additive,
antimultiplicative, involutive star operation), a star-closed non-unital subsemiring is itself a
star ring. -/
instance instStarRing {S R : Type*} [NonUnitalNonAssocSemiring R] [StarRing R] [SetLike S R]
[NonUnitalSubsemiringClass S R] [StarMemClass S R] (s : S) : StarRing s :=
{ StarMemClass.instStarSemigroup s, StarMemClass.instStarAddMonoid s with }
{ StarMemClass.instStarMul s, StarMemClass.instStarAddMonoid s with }

/-- In a star `R`-module (i.e., `star (r • m) = (star r) • m`) any star-closed subset which is also
closed under the scalar action by `R` is itself a star `R`-module. -/
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Star/Pi.lean
Original file line number Diff line number Diff line change
Expand Up @@ -43,7 +43,7 @@ instance [∀ i, Star (f i)] [∀ i, TrivialStar (f i)] : TrivialStar (∀ i, f
instance [∀ i, InvolutiveStar (f i)] : InvolutiveStar (∀ i, f i)
where star_involutive _ := funext fun _ => star_star _

instance [∀ i, Semigroup (f i)] [∀ i, StarSemigroup (f i)] : StarSemigroup (∀ i, f i)
instance [∀ i, Mul (f i)] [∀ i, StarMul (f i)] : StarMul (∀ i, f i)
where star_mul _ _ := funext fun _ => star_mul _ _

instance [∀ i, AddMonoid (f i)] [∀ i, StarAddMonoid (f i)] : StarAddMonoid (∀ i, f i)
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Algebra/Star/Pointwise.lean
Original file line number Diff line number Diff line change
Expand Up @@ -117,7 +117,7 @@ theorem star_singleton {β : Type*} [InvolutiveStar β] (x : β) : ({x} : Set β
rw [mem_star, mem_singleton_iff, mem_singleton_iff, star_eq_iff_star_eq, eq_comm]
#align set.star_singleton Set.star_singleton

protected theorem star_mul [Monoid α] [StarSemigroup α] (s t : Set α) : (s * t)⋆ = t⋆ * s⋆ := by
protected theorem star_mul [Mul α] [StarMul α] (s t : Set α) : (s * t)⋆ = t⋆ * s⋆ := by
simp_rw [← image_star, ← image2_mul, image_image2, image2_image_left, image2_image_right,
star_mul, image2_swap _ s t]
#align set.star_mul Set.star_mul
Expand All @@ -134,7 +134,7 @@ instance [Star α] [TrivialStar α] : TrivialStar (Set α) where
ext1
simp [star_trivial]

protected theorem star_inv [Group α] [StarSemigroup α] (s : Set α) : s⁻¹⋆ = s⋆⁻¹ := by
protected theorem star_inv [Group α] [StarMul α] (s : Set α) : s⁻¹⋆ = s⋆⁻¹ := by
ext
simp only [mem_star, mem_inv, star_inv]
#align set.star_inv Set.star_inv
Expand Down