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

[ refactor ] deprecate Algebra.Structures.IsGroup.{uniqueˡ-⁻¹|uniqueʳ-⁻¹} with knock-ons for Algebra.Module.* #2571

Merged
merged 17 commits into from
Mar 31, 2025
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
26 changes: 26 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -51,6 +51,18 @@ Deprecated names
*ₗ-assoc+comm⇒*ₗ-*ᵣ-assoc ↦ *ₗ-assoc∧comm⇒*ₗ-*ᵣ-assoc
```

* In `Algebra.Modules.Structures.IsLeftModule`:
```agda
uniqueˡ‿⁻ᴹ ↦ Algebra.Module.Properties.LeftModule.inverseˡ-uniqueᴹ
uniqueʳ‿⁻ᴹ ↦ Algebra.Module.Properties.LeftModule.inverseʳ-uniqueᴹ
```

* In `Algebra.Modules.Structures.IsRightModule`:
```agda
uniqueˡ‿⁻ᴹ ↦ Algebra.Module.Properties.RightModule.inverseˡ-uniqueᴹ
uniqueʳ‿⁻ᴹ ↦ Algebra.Module.Properties.RightModule.inverseʳ-uniqueᴹ
```

* In `Algebra.Properties.Magma.Divisibility`:
```agda
∣∣-sym ↦ ∥-sym
Expand Down Expand Up @@ -86,6 +98,12 @@ Deprecated names
∣-trans ↦ ∣ʳ-trans
```

* In `Algebra.Structures.Group`:
```agda
uniqueˡ-⁻¹ ↦ Algebra.Properties.Group.inverseˡ-unique
uniqueʳ-⁻¹ ↦ Algebra.Properties.Group.inverseʳ-unique
```

* In `Data.List.Base`:
```agda
and ↦ Data.Bool.ListAction.and
Expand Down Expand Up @@ -113,6 +131,8 @@ Deprecated names
New modules
-----------

* `Algebra.Module.Properties.{Bimodule|LeftModule|RightModule}`.

* `Data.List.Base.{and|or|any|all}` have been lifted out into `Data.Bool.ListAction`.

* `Data.List.Base.{sum|product}` and their properties have been lifted out into `Data.Nat.ListAction` and `Data.Nat.ListAction.Properties`.
Expand Down Expand Up @@ -155,6 +175,12 @@ Additions to existing modules
commutativeRing : CommutativeRing c ℓ → CommutativeRing (a ⊔ c) (a ⊔ ℓ)
```

* In `Algebra.Modules.Properties`:
```agda
inverseˡ-uniqueᴹ : x +ᴹ y ≈ 0ᴹ → x ≈ -ᴹ y
inverseʳ-uniqueᴹ : x +ᴹ y ≈ 0ᴹ → y ≈ -ᴹ x
```

* In `Algebra.Properties.Magma.Divisibility`:
```agda
∣ˡ-respʳ-≈ : _∣ˡ_ Respectsʳ _≈_
Expand Down
1 change: 0 additions & 1 deletion src/Algebra/Module/Bundles.agda
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,6 @@ open import Algebra.Module.Definitions
open import Function.Base using (flip)
open import Level using (Level; _⊔_; suc)
open import Relation.Binary.Core using (Rel)
open import Relation.Nullary using (¬_)
import Relation.Binary.Reasoning.Setoid as ≈-Reasoning

private
Expand Down
8 changes: 3 additions & 5 deletions src/Algebra/Module/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,7 @@

{-# OPTIONS --cubical-compatible --safe #-}

open import Algebra using (CommutativeRing; Involutive)
open import Algebra.Bundles using (CommutativeRing)
open import Algebra.Module.Bundles using (Module)
open import Level using (Level)

Expand All @@ -18,7 +18,5 @@ module Algebra.Module.Properties

open Module mod
open import Algebra.Module.Properties.Semimodule semimodule public
open import Algebra.Properties.Group using (⁻¹-involutive)

-ᴹ-involutive : Involutive _≈ᴹ_ -ᴹ_
-ᴹ-involutive = ⁻¹-involutive +ᴹ-group
open import Algebra.Module.Properties.Bimodule bimodule public
using (inverseˡ-uniqueᴹ; inverseʳ-uniqueᴹ; -ᴹ-involutive)
21 changes: 21 additions & 0 deletions src/Algebra/Module/Properties/Bimodule.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,21 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Properties of bimodules.
------------------------------------------------------------------------

{-# OPTIONS --cubical-compatible --safe #-}

open import Algebra.Bundles using (Ring)
open import Algebra.Module.Bundles using (Bimodule)
open import Level using (Level)

module Algebra.Module.Properties.Bimodule
{r ℓr s ℓs m ℓm : Level}
{ringR : Ring r ℓr} {ringS : Ring s ℓs}
(mod : Bimodule ringR ringS m ℓm)
where

open Bimodule mod
open import Algebra.Module.Properties.LeftModule leftModule public
using (inverseˡ-uniqueᴹ; inverseʳ-uniqueᴹ; -ᴹ-involutive)
25 changes: 25 additions & 0 deletions src/Algebra/Module/Properties/LeftModule.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,25 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Properties of left modules.
------------------------------------------------------------------------

{-# OPTIONS --cubical-compatible --safe #-}

open import Algebra.Bundles using (Ring)
open import Algebra.Module.Bundles using (LeftModule)
open import Level using (Level)

module Algebra.Module.Properties.LeftModule
{r ℓr m ℓm : Level}
{ring : Ring r ℓr}
(mod : LeftModule ring m ℓm)
where

open Ring ring
open LeftModule mod
open import Algebra.Properties.AbelianGroup +ᴹ-abelianGroup public
using ()
renaming (inverseˡ-unique to inverseˡ-uniqueᴹ
; inverseʳ-unique to inverseʳ-uniqueᴹ
; ⁻¹-involutive to -ᴹ-involutive)
25 changes: 25 additions & 0 deletions src/Algebra/Module/Properties/RightModule.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,25 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Properties of right modules.
------------------------------------------------------------------------

{-# OPTIONS --cubical-compatible --safe #-}

open import Algebra.Bundles using (Ring)
open import Algebra.Module.Bundles using (RightModule)
open import Level using (Level)

module Algebra.Module.Properties.RightModule
{r ℓr m ℓm : Level}
{ring : Ring r ℓr}
(mod : RightModule ring m ℓm)
where

open Ring ring
open RightModule mod
open import Algebra.Properties.AbelianGroup +ᴹ-abelianGroup public
using ()
renaming (inverseˡ-unique to inverseˡ-uniqueᴹ
; inverseʳ-unique to inverseʳ-uniqueᴹ
; ⁻¹-involutive to -ᴹ-involutive)
4 changes: 2 additions & 2 deletions src/Algebra/Module/Properties/Semimodule.agda
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,7 @@

{-# OPTIONS --cubical-compatible --safe #-}

open import Algebra using (CommutativeSemiring)
open import Algebra.Bundles using (CommutativeSemiring)
open import Algebra.Module.Bundles using (Semimodule)
open import Level using (Level)

Expand All @@ -18,8 +18,8 @@ module Algebra.Module.Properties.Semimodule

open CommutativeSemiring semiring
open Semimodule semimod
open import Relation.Nullary.Negation using (contraposition)
open import Relation.Binary.Reasoning.Setoid ≈ᴹ-setoid
open import Relation.Nullary.Negation using (contraposition)

x≈0⇒x*y≈0 : ∀ {x y} → x ≈ 0# → x *ₗ y ≈ᴹ 0ᴹ
x≈0⇒x*y≈0 {x} {y} x≈0 = begin
Expand Down
26 changes: 22 additions & 4 deletions src/Algebra/Module/Structures.agda
Original file line number Diff line number Diff line change
Expand Up @@ -6,10 +6,6 @@

{-# OPTIONS --cubical-compatible --safe #-}

open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.Structures using (IsEquivalence)

module Algebra.Module.Structures where

open import Algebra.Bundles
Expand All @@ -23,6 +19,10 @@ open import Algebra.Module.Definitions
open import Algebra.Structures using (IsCommutativeMonoid; IsAbelianGroup)
open import Data.Product.Base using (_,_; proj₁; proj₂)
open import Level using (Level; _⊔_)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.Structures using (IsEquivalence)


private
variable
Expand Down Expand Up @@ -216,6 +216,15 @@ module _ (ring : Ring r ℓr)
; uniqueˡ-⁻¹ to uniqueˡ‿-ᴹ
; uniqueʳ-⁻¹ to uniqueʳ‿-ᴹ
)
{-# WARNING_ON_USAGE uniqueˡ‿-ᴹ
"Warning: uniqueˡ‿-ᴹ was deprecated in v2.3.
Please use Algebra.Module.Properties.LeftModule.inverseˡ-uniqueᴹ instead."
#-}
{-# WARNING_ON_USAGE uniqueʳ‿-ᴹ
"Warning: uniqueʳ‿-ᴹ was deprecated in v2.3.
Please use Algebra.Module.Properties.LeftModule.inverseʳ-uniqueᴹ instead."
#-}


record IsRightModule (*ᵣ : Opᵣ R M) : Set (r ⊔ m ⊔ ℓr ⊔ ℓm) where
open Defs ≈ᴹ
Expand Down Expand Up @@ -244,6 +253,15 @@ module _ (ring : Ring r ℓr)
; uniqueˡ-⁻¹ to uniqueˡ‿-ᴹ
; uniqueʳ-⁻¹ to uniqueʳ‿-ᴹ
)
{-# WARNING_ON_USAGE uniqueˡ‿-ᴹ
"Warning: uniqueˡ‿-ᴹ was deprecated in v2.3.
Please use Algebra.Module.Properties.RightModule.inverseˡ-uniqueᴹ instead."
#-}
{-# WARNING_ON_USAGE uniqueʳ‿-ᴹ
"Warning: uniqueʳ‿-ᴹ was deprecated in v2.3.
Please use Algebra.Module.Properties.RightModule.inverseʳ-uniqueᴹ instead."
#-}


module _ (R-ring : Ring r ℓr) (S-ring : Ring s ℓs)
(≈ᴹ : Rel {m} M ℓm) (+ᴹ : Op₂ M) (0ᴹ : M) (-ᴹ : Op₁ M)
Expand Down
1 change: 0 additions & 1 deletion src/Algebra/Morphism.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,7 +10,6 @@ module Algebra.Morphism where

import Algebra.Morphism.Definitions as MorphismDefinitions
open import Algebra
import Algebra.Properties.Group as GroupP
open import Function.Base
open import Level
open import Relation.Binary.Core using (Rel; _Preserves_⟶_)
Expand Down
8 changes: 8 additions & 0 deletions src/Algebra/Structures.agda
Original file line number Diff line number Diff line change
Expand Up @@ -323,10 +323,18 @@ record IsGroup (_∙_ : Op₂ A) (ε : A) (_⁻¹ : Op₁ A) : Set (a ⊔ ℓ) w
uniqueˡ-⁻¹ : ∀ x y → (x ∙ y) ≈ ε → x ≈ (y ⁻¹)
uniqueˡ-⁻¹ = Consequences.assoc∧id∧invʳ⇒invˡ-unique
setoid ∙-cong assoc identity inverseʳ
{-# WARNING_ON_USAGE uniqueˡ-⁻¹
"Warning: uniqueˡ-⁻¹ was deprecated in v2.3.
Please use Algebra.Properties.Group.inverseˡ-unique instead. "
#-}

uniqueʳ-⁻¹ : ∀ x y → (x ∙ y) ≈ ε → y ≈ (x ⁻¹)
uniqueʳ-⁻¹ = Consequences.assoc∧id∧invˡ⇒invʳ-unique
setoid ∙-cong assoc identity inverseˡ
{-# WARNING_ON_USAGE uniqueʳ-⁻¹
"Warning: uniqueʳ-⁻¹ was deprecated in v2.3.
Please use Algebra.Properties.Group.inverseʳ-unique instead. "
#-}

isInvertibleMagma : IsInvertibleMagma _∙_ ε _⁻¹
isInvertibleMagma = record
Expand Down