Skip to content

Commit 1c9399b

Browse files
Deprecate Algebra.Operations.Semiring (#1377)
1 parent 8688f42 commit 1c9399b

File tree

8 files changed

+119
-68
lines changed

8 files changed

+119
-68
lines changed

CHANGELOG.md

Lines changed: 10 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -95,6 +95,9 @@ Deprecated modules
9595
The properties of summation in `Algebra.Properties.CommutativeMonoid` have likewise
9696
been deprecated and moved to `Algebra.Properties.CommutativeMonoid.Summation`.
9797

98+
* The module `Algebra.Operations.Semiring` has been deprecated. The contents has
99+
been moved to `Algebra.Properties.Semiring.(Multiplication/Exponentiation)`.
100+
98101
Deprecated names
99102
----------------
100103

@@ -179,7 +182,7 @@ New modules
179182
Algebra.Properties.CommutativeSemigroup.Divisibility
180183
```
181184

182-
* Generic summations over algebraic structures
185+
* Generic summation over algebraic structures
183186
```
184187
Algebra.Properties.Monoid.Summation
185188
Algebra.Properties.CommutativeMonoid.Summation
@@ -188,6 +191,12 @@ New modules
188191
* Generic multiplication over algebraic structures
189192
```
190193
Algebra.Properties.Monoid.Multiplication
194+
Algebra.Properties.Semiring.Multiplication
195+
```
196+
197+
* Generic exponentiation over algebraic structures
198+
```
199+
Algebra.Properties.Semiring.Exponentiation
191200
```
192201

193202
* Setoid equality over vectors:

src/Algebra/Operations/Semiring.agda

Lines changed: 11 additions & 59 deletions
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,7 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Some defined operations (multiplication by natural number and
5-
-- exponentiation)
4+
-- This module is DEPRECATED.
65
------------------------------------------------------------------------
76

87
{-# OPTIONS --without-K --safe #-}
@@ -12,68 +11,21 @@
1211
{-# OPTIONS --warn=noUserWarning #-}
1312

1413
open import Algebra
14+
import Algebra.Operations.CommutativeMonoid as MonoidOperations
1515

1616
module Algebra.Operations.Semiring {s₁ s₂} (S : Semiring s₁ s₂) where
1717

18-
import Algebra.Operations.CommutativeMonoid as MonoidOperations
19-
open import Data.Nat.Base
20-
using (zero; suc; ℕ) renaming (_+_ to _ℕ+_; _*_ to _ℕ*_)
21-
open import Data.Product using (module Σ)
22-
open import Function.Base using (_$_)
23-
open import Relation.Binary
24-
open import Relation.Binary.PropositionalEquality as P using (_≡_)
18+
{-# WARNING_ON_IMPORT
19+
"Algebra.Operations.Semiring was deprecated in v1.5.
20+
Use Algebra.Properties.Semiring.(Multiplication/Exponentiation) instead."
21+
#-}
2522

26-
open Semiring S renaming (zero to *-zero)
27-
open import Relation.Binary.Reasoning.Setoid setoid
23+
open Semiring S
2824

2925
------------------------------------------------------------------------
30-
-- Operations
26+
-- Re-exports
3127

32-
-- Re-export all monoid operations and proofs
3328
open MonoidOperations +-commutativeMonoid public
34-
35-
-- Exponentiation.
36-
infixr 9 _^_
37-
_^_ : Carrier Carrier
38-
x ^ zero = 1#
39-
x ^ suc n = x * x ^ n
40-
41-
------------------------------------------------------------------------
42-
-- Properties of _×_
43-
44-
-- _× 1# is homomorphic with respect to _ℕ*_/_*_.
45-
46-
×1-homo-* : m n (m ℕ* n) × 1# ≈ (m × 1#) * (n × 1#)
47-
×1-homo-* 0 n = begin
48-
0# ≈⟨ sym (Σ.proj₁ *-zero (n × 1#)) ⟩
49-
0# * (n × 1#) ∎
50-
×1-homo-* (suc m) n = begin
51-
(n ℕ+ m ℕ* n) × 1# ≈⟨ ×-homo-+ 1# n (m ℕ* n) ⟩
52-
n × 1# + (m ℕ* n) × 1# ≈⟨ +-cong refl (×1-homo-* m n) ⟩
53-
n × 1# + (m × 1#) * (n × 1#) ≈⟨ sym (+-cong (*-identityˡ _) refl) ⟩
54-
1# * (n × 1#) + (m × 1#) * (n × 1#) ≈⟨ sym (distribʳ (n × 1#) 1# (m × 1#)) ⟩
55-
(1# + m × 1#) * (n × 1#) ∎
56-
57-
------------------------------------------------------------------------
58-
-- Properties of _×′_
59-
60-
-- _×′ 1# is homomorphic with respect to _ℕ*_/_*_.
61-
62-
×′1-homo-* : m n (m ℕ* n) ×′ 1# ≈ (m ×′ 1#) * (n ×′ 1#)
63-
×′1-homo-* m n = begin
64-
(m ℕ* n) ×′ 1# ≈⟨ sym $ ×≈×′ (m ℕ* n) 1# ⟩
65-
(m ℕ* n) × 1# ≈⟨ ×1-homo-* m n ⟩
66-
(m × 1#) * (n × 1#) ≈⟨ *-cong (×≈×′ m 1#) (×≈×′ n 1#) ⟩
67-
(m ×′ 1#) * (n ×′ 1#) ∎
68-
69-
------------------------------------------------------------------------
70-
-- Properties of _^_
71-
72-
-- _^_ preserves equality.
73-
74-
^-congˡ : n (_^ n) Preserves _≈_ ⟶ _≈_
75-
^-congˡ zero x≈y = refl
76-
^-congˡ (suc n) x≈y = *-cong x≈y (^-congˡ n x≈y)
77-
78-
^-cong : _^_ Preserves₂ _≈_ ⟶ _≡_ ⟶ _≈_
79-
^-cong {v = n} x≈y P.refl = ^-congˡ n x≈y
29+
open import Algebra.Properties.Semiring.Exponentiation S public
30+
open import Algebra.Properties.Semiring.Multiplication S public
31+
using (×1-homo-*; ×′1-homo-*)

src/Algebra/Properties/Monoid/Multiplication.agda

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ open Monoid M
1818
renaming
1919
( _∙_ to _+_
2020
; ∙-cong to +-cong
21-
; ∙-congˡ to -congˡ
21+
; ∙-congˡ to +-congˡ
2222
; identityˡ to +-identityˡ
2323
; identityʳ to +-identityʳ
2424
; assoc to +-assoc
@@ -69,7 +69,7 @@ suc n ×′ x = x + n ×′ x
6969
n .{{_ : NonZero n}} n × c ≈ c
7070
×-idem {c} idem (suc zero) = +-identityʳ c
7171
×-idem {c} idem (suc (suc n)) = begin
72-
c + (suc n × c) ≈⟨ -congˡ (×-idem idem (suc n) ) ⟩
72+
c + (suc n × c) ≈⟨ +-congˡ (×-idem idem (suc n) ) ⟩
7373
c + c ≈⟨ idem ⟩
7474
c ∎
7575

@@ -86,7 +86,7 @@ suc n ×′ x = x + n ×′ x
8686
×≈×′ : n x n × x ≈ n ×′ x
8787
×≈×′ 0 x = refl
8888
×≈×′ (suc n) x = begin
89-
x + n × x ≈⟨ +-cong refl (×≈×′ n x) ⟩
89+
x + n × x ≈⟨ +-congˡ (×≈×′ n x) ⟩
9090
x + n ×′ x ≈⟨ sym (1+×′ n x) ⟩
9191
suc n ×′ x ∎
9292

Lines changed: 42 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,42 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Exponentiation defined over a semiring as repeated multiplication
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --without-K --safe #-}
8+
9+
open import Algebra
10+
11+
module Algebra.Properties.Semiring.Exponentiation
12+
{a ℓ} (S : Semiring a ℓ) where
13+
14+
open import Data.Nat.Base using (zero; suc; ℕ)
15+
open import Relation.Binary
16+
open import Relation.Binary.PropositionalEquality as P using (_≡_)
17+
18+
open Semiring S renaming (zero to *-zero)
19+
open import Relation.Binary.Reasoning.Setoid setoid
20+
21+
------------------------------------------------------------------------
22+
-- Operations
23+
24+
-- Standard exponentiation.
25+
26+
infixr 9 _^_
27+
28+
_^_ : Carrier Carrier
29+
x ^ zero = 1#
30+
x ^ suc n = x * x ^ n
31+
32+
------------------------------------------------------------------------
33+
-- Properties of _^_
34+
35+
-- _^_ preserves equality.
36+
37+
^-congˡ : n (_^ n) Preserves _≈_ ⟶ _≈_
38+
^-congˡ zero x≈y = refl
39+
^-congˡ (suc n) x≈y = *-cong x≈y (^-congˡ n x≈y)
40+
41+
^-cong : _^_ Preserves₂ _≈_ ⟶ _≡_ ⟶ _≈_
42+
^-cong {v = n} x≈y P.refl = ^-congˡ n x≈y
Lines changed: 48 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,48 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Multiplication by a natural number over a semiring
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --without-K --safe #-}
8+
9+
open import Algebra
10+
import Algebra.Properties.Monoid.Multiplication as MonoidMultiplication
11+
open import Data.Nat.Base as ℕ using (zero; suc)
12+
13+
module Algebra.Properties.Semiring.Multiplication
14+
{a ℓ} (S : Semiring a ℓ) where
15+
16+
open Semiring S renaming (zero to *-zero)
17+
open import Relation.Binary.Reasoning.Setoid setoid
18+
19+
------------------------------------------------------------------------
20+
-- Re-export definition from the monoid
21+
22+
open MonoidMultiplication +-monoid public
23+
24+
------------------------------------------------------------------------
25+
-- Properties of _×_
26+
27+
-- (_× 1#) is homomorphic with respect to _ℕ.*_/_*_.
28+
29+
×1-homo-* : m n (m ℕ.* n) × 1# ≈ (m × 1#) * (n × 1#)
30+
×1-homo-* 0 n = sym (zeroˡ (n × 1#))
31+
×1-homo-* (suc m) n = begin
32+
(n ℕ.+ m ℕ.* n) × 1# ≈⟨ ×-homo-+ 1# n (m ℕ.* n) ⟩
33+
n × 1# + (m ℕ.* n) × 1# ≈⟨ +-congˡ (×1-homo-* m n) ⟩
34+
n × 1# + (m × 1#) * (n × 1#) ≈˘⟨ +-congʳ (*-identityˡ _) ⟩
35+
1# * (n × 1#) + (m × 1#) * (n × 1#) ≈˘⟨ distribʳ (n × 1#) 1# (m × 1#) ⟩
36+
(1# + m × 1#) * (n × 1#) ∎
37+
38+
------------------------------------------------------------------------
39+
-- Properties of _×′_
40+
41+
-- (_×′ 1#) is homomorphic with respect to _ℕ.*_/_*_.
42+
43+
×′1-homo-* : m n (m ℕ.* n) ×′ 1# ≈ (m ×′ 1#) * (n ×′ 1#)
44+
×′1-homo-* m n = begin
45+
(m ℕ.* n) ×′ 1# ≈˘⟨ ×≈×′ (m ℕ.* n) 1# ⟩
46+
(m ℕ.* n) × 1# ≈⟨ ×1-homo-* m n ⟩
47+
(m × 1#) * (n × 1#) ≈⟨ *-cong (×≈×′ m 1#) (×≈×′ n 1#) ⟩
48+
(m ×′ 1#) * (n ×′ 1#) ∎

src/Algebra/Solver/Ring.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -37,7 +37,7 @@ open AlmostCommutativeRing R
3737
open import Algebra.Definitions _≈_
3838
open import Algebra.Morphism
3939
open _-Raw-AlmostCommutative⟶_ morphism renaming (⟦_⟧ to ⟦_⟧′)
40-
open import Algebra.Operations.Semiring semiring
40+
open import Algebra.Properties.Semiring.Exponentiation semiring
4141

4242
open import Relation.Binary
4343
open import Relation.Nullary using (yes; no)

src/Algebra/Solver/Ring/NaturalCoefficients.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -8,13 +8,13 @@
88
{-# OPTIONS --without-K --safe #-}
99

1010
open import Algebra
11-
import Algebra.Operations.Semiring as SemiringOps
11+
import Algebra.Properties.Semiring.Multiplication as SemiringMultiplication
1212
open import Data.Maybe.Base using (Maybe; just; nothing; map)
1313

1414
module Algebra.Solver.Ring.NaturalCoefficients
1515
{r₁ r₂} (R : CommutativeSemiring r₁ r₂)
1616
(open CommutativeSemiring R)
17-
(open SemiringOps semiring)
17+
(open SemiringMultiplication semiring)
1818
(dec : m n Maybe (m × 1# ≈ n × 1#)) where
1919

2020
import Algebra.Solver.Ring

src/Algebra/Solver/Ring/NaturalCoefficients/Default.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -16,14 +16,14 @@ open import Algebra
1616
module Algebra.Solver.Ring.NaturalCoefficients.Default
1717
{r₁ r₂} (R : CommutativeSemiring r₁ r₂) where
1818

19-
import Algebra.Operations.Semiring as SemiringOps
19+
import Algebra.Properties.Semiring.Multiplication as SemiringMultiplication
2020
open import Data.Maybe.Base using (Maybe; map)
2121
open import Data.Nat using (_≟_)
2222
open import Relation.Binary.Consequences using (dec⟶weaklyDec)
2323
import Relation.Binary.PropositionalEquality as P
2424

2525
open CommutativeSemiring R
26-
open SemiringOps semiring
26+
open SemiringMultiplication semiring
2727

2828
private
2929
dec : m n Maybe (m × 1# ≈ n × 1#)

0 commit comments

Comments
 (0)