Skip to content

Commit 6d6c6df

Browse files
authored
Add coincidence law to modules (#1898)
1 parent 2378d52 commit 6d6c6df

File tree

7 files changed

+50
-9
lines changed

7 files changed

+50
-9
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -1843,6 +1843,11 @@ New modules
18431843
Algebra.Module.Core
18441844
```
18451845
1846+
* Definitions for module-like algebraic structures with left- and right- multiplication over the same scalars:
1847+
```
1848+
Algebra.Module.Definitions.Bi.Simultaneous
1849+
```
1850+
18461851
* Constructive algebraic structures with apartness relations:
18471852
```
18481853
Algebra.Apartness

src/Algebra/Module/Construct/DirectProduct.agda

+2
Original file line numberDiff line numberDiff line change
@@ -97,6 +97,7 @@ semimodule M N = record
9797
{ isSemimodule = record
9898
{ isBisemimodule = Bisemimodule.isBisemimodule
9999
(bisemimodule M.bisemimodule N.bisemimodule)
100+
; *ₗ-*ᵣ-coincident = λ x m M.*ₗ-*ᵣ-coincident x (proj₁ m) , N.*ₗ-*ᵣ-coincident x (proj₂ m)
100101
}
101102
} where module M = Semimodule M; module N = Semimodule N
102103

@@ -155,5 +156,6 @@ bimodule M N = record
155156
⟨module⟩ M N = record
156157
{ isModule = record
157158
{ isBimodule = Bimodule.isBimodule (bimodule M.bimodule N.bimodule)
159+
; *ₗ-*ᵣ-coincident = λ x m M.*ₗ-*ᵣ-coincident x (proj₁ m) , N.*ₗ-*ᵣ-coincident x (proj₂ m)
158160
}
159161
} where module M = Module M; module N = Module N

src/Algebra/Module/Construct/TensorUnit.agda

+2
Original file line numberDiff line numberDiff line change
@@ -77,6 +77,7 @@ semimodule : {R : CommutativeSemiring c ℓ} → Semimodule R c ℓ
7777
semimodule {R = commutativeSemiring} = record
7878
{ isSemimodule = record
7979
{ isBisemimodule = Bisemimodule.isBisemimodule bisemimodule
80+
; *ₗ-*ᵣ-coincident = *-comm
8081
}
8182
} where open CommutativeSemiring commutativeSemiring
8283

@@ -113,5 +114,6 @@ bimodule {R = ring} = record
113114
⟨module⟩ {R = commutativeRing} = record
114115
{ isModule = record
115116
{ isBimodule = Bimodule.isBimodule bimodule
117+
; *ₗ-*ᵣ-coincident = *-comm
116118
}
117119
} where open CommutativeRing commutativeRing

src/Algebra/Module/Definitions.agda

+2
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,9 @@ module Algebra.Module.Definitions where
1212
import Algebra.Module.Definitions.Left as L
1313
import Algebra.Module.Definitions.Right as R
1414
import Algebra.Module.Definitions.Bi as B
15+
import Algebra.Module.Definitions.Bi.Simultaneous as BS
1516

1617
module LeftDefs = L
1718
module RightDefs = R
1819
module BiDefs = B
20+
module SimultaneousBiDefs = BS
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,18 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Properties connecting left-scaling and right-scaling over the same scalars
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Relation.Binary
10+
11+
module Algebra.Module.Definitions.Bi.Simultaneous
12+
{a b ℓb} (A : Set a) {B : Set b} (_≈_ : Rel B ℓb)
13+
where
14+
15+
open import Algebra.Module.Core
16+
17+
Coincident : Opₗ A B Opᵣ A B Set _
18+
Coincident _∙ₗ_ _∙ᵣ_ = x m (x ∙ₗ m) ≈ (m ∙ᵣ x)

src/Algebra/Module/Structures.agda

+11-7
Original file line numberDiff line numberDiff line change
@@ -166,14 +166,16 @@ module _ (commutativeSemiring : CommutativeSemiring r ℓr)
166166
open CommutativeSemiring commutativeSemiring renaming (Carrier to R)
167167

168168
-- An R-semimodule is an R-R-bisemimodule where R is commutative.
169-
-- This means that *ₗ and *ᵣ coincide up to mathematical equality,
170-
-- though it may be that they do not coincide up to definitional
171-
-- equality.
169+
-- We enforce that *ₗ and *ᵣ coincide up to mathematical equality, though it
170+
-- may be that they do not coincide up to definitional equality.
171+
172+
open SimultaneousBiDefs R ≈ᴹ
172173

173174
record IsSemimodule (*ₗ : Opₗ R M) (*ᵣ : Opᵣ R M)
174175
: Set (r ⊔ m ⊔ ℓr ⊔ ℓm) where
175176
field
176177
isBisemimodule : IsBisemimodule semiring semiring ≈ᴹ +ᴹ 0ᴹ *ₗ *ᵣ
178+
*ₗ-*ᵣ-coincident : Coincident *ₗ *ᵣ
177179

178180
open IsBisemimodule isBisemimodule public
179181

@@ -282,15 +284,17 @@ module _ (commutativeRing : CommutativeRing r ℓr)
282284
open CommutativeRing commutativeRing renaming (Carrier to R)
283285

284286
-- An R-module is an R-R-bimodule where R is commutative.
285-
-- This means that *ₗ and *ᵣ coincide up to mathematical equality,
286-
-- though it may be that they do not coincide up to definitional
287-
-- equality.
287+
-- We enforce that *ₗ and *ᵣ coincide up to mathematical equality, though it
288+
-- may be that they do not coincide up to definitional equality.
289+
290+
open SimultaneousBiDefs R ≈ᴹ
288291

289292
record IsModule (*ₗ : Opₗ R M) (*ᵣ : Opᵣ R M) : Set (r ⊔ m ⊔ ℓr ⊔ ℓm) where
290293
field
291294
isBimodule : IsBimodule ring ring ≈ᴹ +ᴹ 0ᴹ -ᴹ *ₗ *ᵣ
295+
*ₗ-*ᵣ-coincident : Coincident *ₗ *ᵣ
292296

293297
open IsBimodule isBimodule public
294298

295299
isSemimodule : IsSemimodule commutativeSemiring ≈ᴹ +ᴹ 0ᴹ *ₗ *ᵣ
296-
isSemimodule = record { isBisemimodule = isBisemimodule }
300+
isSemimodule = record { isBisemimodule = isBisemimodule; *ₗ-*ᵣ-coincident = *ₗ-*ᵣ-coincident }

src/Algebra/Module/Structures/Biased.agda

+10-2
Original file line numberDiff line numberDiff line change
@@ -57,7 +57,10 @@ module _ (commutativeSemiring : CommutativeSemiring r ℓr) where
5757
}
5858

5959
isSemimodule : IsSemimodule commutativeSemiring ≈ᴹ +ᴹ 0ᴹ *ₗ (flip *ₗ)
60-
isSemimodule = record { isBisemimodule = isBisemimodule }
60+
isSemimodule = record
61+
{ isBisemimodule = isBisemimodule
62+
; *ₗ-*ᵣ-coincident = λ _ _ ≈ᴹ-refl
63+
}
6164

6265
-- Similarly, a right semimodule over a commutative semiring
6366
-- is already a semimodule.
@@ -88,7 +91,10 @@ module _ (commutativeSemiring : CommutativeSemiring r ℓr) where
8891
}
8992

9093
isSemimodule : IsSemimodule commutativeSemiring ≈ᴹ +ᴹ 0ᴹ (flip *ᵣ) *ᵣ
91-
isSemimodule = record { isBisemimodule = isBisemimodule }
94+
isSemimodule = record
95+
{ isBisemimodule = isBisemimodule
96+
; *ₗ-*ᵣ-coincident = λ _ _ ≈ᴹ-refl
97+
}
9298

9399

94100
module _ (commutativeRing : CommutativeRing r ℓr) where
@@ -113,6 +119,7 @@ module _ (commutativeRing : CommutativeRing r ℓr) where
113119
; -ᴹ‿cong = -ᴹ‿cong
114120
; -ᴹ‿inverse = -ᴹ‿inverse
115121
}
122+
; *ₗ-*ᵣ-coincident = λ _ _ ≈ᴹ-refl
116123
}
117124

118125
-- Similarly, a right module over a commutative ring is already a module.
@@ -134,4 +141,5 @@ module _ (commutativeRing : CommutativeRing r ℓr) where
134141
; -ᴹ‿cong = -ᴹ‿cong
135142
; -ᴹ‿inverse = -ᴹ‿inverse
136143
}
144+
; *ₗ-*ᵣ-coincident = λ _ _ ≈ᴹ-refl
137145
}

0 commit comments

Comments
 (0)