Skip to content

Commit 01e20a8

Browse files
JacquesCaretteandreasabel
authored andcommitted
Add a consequence of definition straight to IsCancellativeCommutativeSemiring (#2370)
* Closes #2315 * whitespace
1 parent 1d0231b commit 01e20a8

File tree

2 files changed

+10
-0
lines changed

2 files changed

+10
-0
lines changed

CHANGELOG.md

+6
Original file line numberDiff line numberDiff line change
@@ -249,6 +249,12 @@ Additions to existing modules
249249
x \\ y = (x ⁻¹) ∙ y
250250
```
251251

252+
* In `Algebra.Structures.IsCancellativeCommutativeSemiring` add the
253+
extra property as an exposed definition:
254+
```agda
255+
*-cancelʳ-nonZero : AlmostRightCancellative 0# *
256+
```
257+
252258
* In `Data.Container.Indexed.Core`:
253259
```agda
254260
Subtrees o c = (r : Response c) → X (next c r)

src/Algebra/Structures.agda

+4
Original file line numberDiff line numberDiff line change
@@ -573,6 +573,10 @@ record IsCancellativeCommutativeSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a
573573

574574
open IsCommutativeSemiring isCommutativeSemiring public
575575

576+
*-cancelʳ-nonZero : AlmostRightCancellative 0# *
577+
*-cancelʳ-nonZero = Consequences.comm∧almostCancelˡ⇒almostCancelʳ setoid
578+
*-comm *-cancelˡ-nonZero
579+
576580
record IsIdempotentSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
577581
field
578582
isSemiring : IsSemiring + * 0# 1#

0 commit comments

Comments
 (0)