Skip to content

Commit 147d31d

Browse files
Deprecate Algebra.Operations.CommutativeMonoid (#1351)
1 parent be9ac3a commit 147d31d

File tree

9 files changed

+459
-119
lines changed

9 files changed

+459
-119
lines changed

CHANGELOG.md

+24
Original file line numberDiff line numberDiff line change
@@ -63,6 +63,14 @@ Deprecated modules
6363
complete. The new definitions are parameterised by raw bundles instead of bundles
6464
meaning they are much more flexible to work with.
6565

66+
* The module `Algebra.Operations.CommutativeMonoid` has been deprecated. The definition
67+
of multiplication and the associated properties have been moved to
68+
`Algebra.Properties.CommutativeMonoid.Multiplication`. The definition of summation
69+
which was defined over the deprecated `Data.Table` has been redefined in terms of
70+
`Data.Vec.Functional` and been moved to `Algbra.Properties.CommutativeMonoid.Summation`.
71+
The properties of summation in `Algebra.Properties.CommutativeMonoid` have likewise
72+
been deprecated and moved to `Algebra.Properties.CommutativeMonoid.Summation`.
73+
6674
Deprecated names
6775
----------------
6876

@@ -125,6 +133,22 @@ New modules
125133
Algebra.Properties.CommutativeSemigroup.Divisibility
126134
```
127135

136+
* Generic summations over algebraic structures
137+
```
138+
Algebra.Properties.Monoid.Summation
139+
Algebra.Properties.CommutativeMonoid.Summation
140+
```
141+
142+
* Generic multiplication over algebraic structures
143+
```
144+
Algebra.Properties.Monoid.Multiplication
145+
```
146+
147+
* Setoid equality over vectors:
148+
```
149+
Data.Vec.Functional.Relation.Binary.Equality.Setoid
150+
```
151+
128152
* Heterogeneous relation characterising a list as an infix segment of another:
129153
```
130154
Data.List.Relation.Binary.Infix.Heterogeneous
+12-87
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 #-}
@@ -11,51 +10,26 @@
1110
{-# OPTIONS --warn=noUserWarning #-}
1211

1312
open import Algebra
13+
open import Data.List.Base as List using (List; []; _∷_; _++_)
14+
open import Data.Fin.Base using (Fin; zero)
15+
open import Data.Table.Base as Table using (Table)
16+
open import Function using (_∘_)
17+
open import Relation.Binary.PropositionalEquality as P using (_≡_)
1418

1519
module Algebra.Operations.CommutativeMonoid
1620
{s₁ s₂} (CM : CommutativeMonoid s₁ s₂)
1721
where
1822

19-
open import Data.Nat.Base using (ℕ; zero; suc)
20-
renaming (_+_ to _ℕ+_; _*_ to _ℕ*_)
21-
open import Data.List.Base as List using (List; []; _∷_; _++_)
22-
open import Data.Fin.Base using (Fin; zero)
23-
open import Data.Product using (proj₁; proj₂)
24-
open import Data.Table.Base as Table using (Table)
25-
open import Function using (_∘_; _⟨_⟩_)
26-
open import Relation.Binary
27-
open import Relation.Binary.PropositionalEquality as P using (_≡_)
23+
{-# WARNING_ON_IMPORT
24+
"Algebra.Operations.CommutativeMonoid was deprecated in v1.5.
25+
Use Algebra.Properties.CommutativeMonoid.(Summation/Multiplication) instead."
26+
#-}
2827

2928
open CommutativeMonoid CM
3029
renaming
3130
( _∙_ to _+_
32-
; ∙-cong to +-cong
33-
; identityˡ to +-identityˡ
34-
; identityʳ to +-identityʳ
35-
; assoc to +-assoc
36-
; comm to +-comm
3731
; ε to 0#
3832
)
39-
open import Relation.Binary.Reasoning.Setoid setoid
40-
41-
------------------------------------------------------------------------
42-
-- Operations
43-
44-
-- Multiplication by natural number.
45-
46-
infixr 8 _×_ _×′_
47-
48-
_×_ : Carrier Carrier
49-
0 × x = 0#
50-
suc n × x = x + n × x
51-
52-
-- A variant that includes a "redundant" case which ensures that `1 × x`
53-
-- is definitionally equal to `x`.
54-
55-
_×′_ : Carrier Carrier
56-
0 ×′ x = 0#
57-
1 ×′ x = x
58-
suc n ×′ x = x + n ×′ x
5933

6034
-- Summation over lists/tables
6135

@@ -75,55 +49,6 @@ sumₜ-syntax _ = sumₜ ∘ Table.tabulate
7549
syntax sumₜ-syntax n (λ i x) = ∑[ i < n ] x
7650

7751
------------------------------------------------------------------------
78-
-- Properties of _×_
79-
80-
×-congʳ : n (n ×_) Preserves _≈_ ⟶ _≈_
81-
×-congʳ 0 x≈x′ = refl
82-
×-congʳ (suc n) x≈x′ = x≈x′ ⟨ +-cong ⟩ ×-congʳ n x≈x′
83-
84-
×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
85-
×-cong {u} P.refl x≈x′ = ×-congʳ u x≈x′
86-
87-
-- _×_ is homomorphic with respect to _ℕ+_/_+_.
88-
89-
×-homo-+ : c m n (m ℕ+ n) × c ≈ m × c + n × c
90-
×-homo-+ c 0 n = sym (+-identityˡ (n × c))
91-
×-homo-+ c (suc m) n = begin
92-
c + (m ℕ+ n) × c ≈⟨ +-cong refl (×-homo-+ c m n) ⟩
93-
c + (m × c + n × c) ≈⟨ sym (+-assoc c (m × c) (n × c)) ⟩
94-
c + m × c + n × c ∎
95-
96-
------------------------------------------------------------------------
97-
-- Properties of _×′_
98-
99-
1+×′ : n x suc n ×′ x ≈ x + n ×′ x
100-
1+×′ 0 x = sym (+-identityʳ x)
101-
1+×′ (suc n) x = refl
102-
103-
-- _×_ and _×′_ are extensionally equal (up to the setoid
104-
-- equivalence).
105-
106-
×≈×′ : n x n × x ≈ n ×′ x
107-
×≈×′ 0 x = begin 0# ∎
108-
×≈×′ (suc n) x = begin
109-
x + n × x ≈⟨ +-cong refl (×≈×′ n x) ⟩
110-
x + n ×′ x ≈⟨ sym (1+×′ n x) ⟩
111-
suc n ×′ x ∎
112-
113-
-- _×′_ is homomorphic with respect to _ℕ+_/_+_.
114-
115-
×′-homo-+ : c m n (m ℕ+ n) ×′ c ≈ m ×′ c + n ×′ c
116-
×′-homo-+ c m n = begin
117-
(m ℕ+ n) ×′ c ≈⟨ sym (×≈×′ (m ℕ+ n) c) ⟩
118-
(m ℕ+ n) × c ≈⟨ ×-homo-+ c m n ⟩
119-
m × c + n × c ≈⟨ +-cong (×≈×′ m c) (×≈×′ n c) ⟩
120-
m ×′ c + n ×′ c ∎
121-
122-
-- _×′_ preserves equality.
52+
-- Operations
12353

124-
×′-cong : _×′_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
125-
×′-cong {n} {_} {x} {y} P.refl x≈y = begin
126-
n ×′ x ≈⟨ sym (×≈×′ n x) ⟩
127-
n × x ≈⟨ ×-congʳ n x≈y ⟩
128-
n × y ≈⟨ ×≈×′ n y ⟩
129-
n ×′ y ∎
54+
open import Algebra.Properties.Monoid.Multiplication monoid public

src/Algebra/Operations/Semiring.agda

+4
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,10 @@
77

88
{-# OPTIONS --without-K --safe #-}
99

10+
-- Disabled to prevent warnings from deprecated
11+
-- Algebra.Operations.CommutativeMonoid
12+
{-# OPTIONS --warn=noUserWarning #-}
13+
1014
open import Algebra
1115

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

src/Algebra/Properties/CommutativeMonoid.agda

+59-31
Original file line numberDiff line numberDiff line change
@@ -52,41 +52,35 @@ open CommutativeMonoid M
5252
open import Algebra.Definitions _≈_
5353
open import Relation.Binary.Reasoning.Setoid setoid
5454

55-
module _ {n} where
56-
open B.Setoid (TE.setoid setoid n) public
57-
using ()
58-
renaming (_≈_ to _≋_)
5955

60-
-- When summing over a function from a finite set, we can pull out any value and move it to the front.
6156

62-
sumₜ-remove : {n} {i : Fin (suc n)} t sumₜ t ≈ lookup t i + sumₜ (remove i t)
63-
sumₜ-remove {_} {zero} t = refl
64-
sumₜ-remove {suc n} {suc i} t′ =
65-
begin
66-
t₀ + ∑t ≈⟨ +-congˡ (sumₜ-remove t) ⟩
67-
t₀ + (tᵢ + ∑t′) ≈⟨ solve 3 (λ x y z x ⊕ (y ⊕ z) ⊜ y ⊕ (x ⊕ z)) refl t₀ tᵢ ∑t′ ⟩
68-
tᵢ + (t₀ + ∑t′) ∎
69-
where
70-
t = tail t′
71-
t₀ = head t′
72-
tᵢ = lookup t i
73-
∑t = sumₜ t
74-
∑t′ = sumₜ (remove i t)
57+
------------------------------------------------------------------------
58+
-- DEPRECATED NAMES
59+
------------------------------------------------------------------------
60+
-- Please use the new names as continuing support for the old names is
61+
-- not guaranteed.
62+
63+
module _ {n} where
64+
open B.Setoid (TE.setoid setoid n) public
65+
using () renaming (_≈_ to _≋_)
7566

76-
-- '_≈_' is a congruence over 'sumTable n'.
67+
-- Version 1.5
7768

7869
sumₜ-cong-≈ : {n} sumₜ {n} Preserves _≋_ ⟶ _≈_
7970
sumₜ-cong-≈ {zero} p = refl
8071
sumₜ-cong-≈ {suc n} p = +-cong (p _) (sumₜ-cong-≈ (p ∘ suc))
81-
82-
-- '_≡_' is a congruence over 'sum n'.
72+
{-# WARNING_ON_USAGE sumₜ-cong-≈
73+
"Warning: sumₜ-cong-≈ was deprecated in v1.5.
74+
Please use sum-cong-≋ from `Algebra.Properties.CommutativeMonoid.Summation` instead."
75+
#-}
8376

8477
sumₜ-cong-≡ : {n} sumₜ {n} Preserves _≗_ ⟶ _≡_
8578
sumₜ-cong-≡ {zero} p = P.refl
8679
sumₜ-cong-≡ {suc n} p = P.cong₂ _+_ (p _) (sumₜ-cong-≡ (p ∘ suc))
87-
88-
-- If addition is idempotent on a particular value 'x', then summing over a
89-
-- nonzero number of copies of 'x' gives back 'x'.
80+
{-# WARNING_ON_USAGE sumₜ-cong-≡
81+
"Warning: sumₜ-cong-≡ was deprecated in v1.5.
82+
Please use sum-cong-≗ from `Algebra.Properties.CommutativeMonoid.Summation` instead."
83+
#-}
9084

9185
sumₜ-idem-replicate : n {x} _+_ IdempotentOn x sumₜ (replicate {n = suc n} x) ≈ x
9286
sumₜ-idem-replicate zero idem = +-identityʳ _
@@ -95,16 +89,38 @@ sumₜ-idem-replicate (suc n) {x} idem = begin
9589
(x + x) + sumₜ (replicate {n = n} x) ≈⟨ +-congʳ idem ⟩
9690
x + sumₜ (replicate {n = n} x) ≈⟨ sumₜ-idem-replicate n idem ⟩
9791
x ∎
98-
99-
-- The sum over the constantly zero function is zero.
92+
{-# WARNING_ON_USAGE sumₜ-idem-replicate
93+
"Warning: sumₜ-idem-replicate was deprecated in v1.5.
94+
Please use sum-replicate-idem from `Algebra.Properties.CommutativeMonoid.Summation` instead."
95+
#-}
10096

10197
sumₜ-zero : n sumₜ (replicate {n = n} 0#) ≈ 0#
10298
sumₜ-zero n = begin
10399
sumₜ (replicate {n = n} 0#) ≈⟨ sym (+-identityˡ _) ⟩
104100
0# + sumₜ (replicate {n = n} 0#) ≈⟨ sumₜ-idem-replicate n (+-identityˡ 0#) ⟩
105101
0# ∎
102+
{-# WARNING_ON_USAGE sumₜ-zero
103+
"Warning: sumₜ-zero was deprecated in v1.5.
104+
Please use sum-replicate-zero from `Algebra.Properties.CommutativeMonoid.Summation` instead."
105+
#-}
106106

107-
-- The '∑' operator distributes over addition.
107+
sumₜ-remove : {n} {i : Fin (suc n)} t sumₜ t ≈ lookup t i + sumₜ (remove i t)
108+
sumₜ-remove {_} {zero} t = refl
109+
sumₜ-remove {suc n} {suc i} t′ =
110+
begin
111+
t₀ + ∑t ≈⟨ +-congˡ (sumₜ-remove t) ⟩
112+
t₀ + (tᵢ + ∑t′) ≈⟨ solve 3 (λ x y z x ⊕ (y ⊕ z) ⊜ y ⊕ (x ⊕ z)) refl t₀ tᵢ ∑t′ ⟩
113+
tᵢ + (t₀ + ∑t′) ∎
114+
where
115+
t = tail t′
116+
t₀ = head t′
117+
tᵢ = lookup t i
118+
∑t = sumₜ t
119+
∑t′ = sumₜ (remove i t)
120+
{-# WARNING_ON_USAGE sumₜ-remove
121+
"Warning: sumₜ-remove was deprecated in v1.5.
122+
Please use sum-remove from `Algebra.Properties.CommutativeMonoid.Summation` instead."
123+
#-}
108124

109125
∑-distrib-+ : n (f g : Fin n Carrier) ∑[ i < n ] (f i + g i) ≈ ∑[ i < n ] f i + ∑[ i < n ] g i
110126
∑-distrib-+ zero f g = sym (+-identityˡ _)
@@ -119,17 +135,21 @@ sumₜ-zero n = begin
119135
∑f = ∑[ i < n ] f (suc i)
120136
∑g = ∑[ i < n ] g (suc i)
121137
∑fg = ∑[ i < n ] (f (suc i) + g (suc i))
122-
123-
-- The '∑' operator commutes with itself.
138+
{-# WARNING_ON_USAGE ∑-distrib-+
139+
"Warning: ∑-distrib-+ was deprecated in v1.5.
140+
Please use ∑-distrib-+ from `Algebra.Properties.CommutativeMonoid.Summation` instead."
141+
#-}
124142

125143
∑-comm : n m (f : Fin n Fin m Carrier) ∑[ i < n ] ∑[ j < m ] f i j ≈ ∑[ j < m ] ∑[ i < n ] f i j
126144
∑-comm zero m f = sym (sumₜ-zero m)
127145
∑-comm (suc n) m f = begin
128146
∑[ j < m ] f zero j + ∑[ i < n ] ∑[ j < m ] f (suc i) j ≈⟨ +-congˡ (∑-comm n m _) ⟩
129147
∑[ j < m ] f zero j + ∑[ j < m ] ∑[ i < n ] f (suc i) j ≈⟨ sym (∑-distrib-+ m _ _) ⟩
130148
∑[ j < m ] (f zero j + ∑[ i < n ] f (suc i) j) ∎
131-
132-
-- Any permutation of a table has the same sum as the original.
149+
{-# WARNING_ON_USAGE ∑-distrib-+
150+
"Warning: ∑-comm was deprecated in v1.5.
151+
Please use ∑-comm from `Algebra.Properties.CommutativeMonoid.Summation` instead."
152+
#-}
133153

134154
sumₜ-permute : {m n} t (π : Permutation m n) sumₜ t ≈ sumₜ (permute π t)
135155
sumₜ-permute {zero} {zero} t π = refl
@@ -145,9 +165,17 @@ sumₜ-permute {suc m} {suc n} t π = begin
145165
where
146166
0i = zero
147167
πt = permute π t
168+
{-# WARNING_ON_USAGE sumₜ-permute
169+
"Warning: sumₜ-permute was deprecated in v1.5.
170+
Please use sum-permute from `Algebra.Properties.CommutativeMonoid.Summation` instead."
171+
#-}
148172

149173
∑-permute : {m n} f (π : Permutation m n) ∑[ i < n ] f i ≈ ∑[ i < m ] f (π ⟨$⟩ʳ i)
150174
∑-permute = sumₜ-permute ∘ tabulate
175+
{-# WARNING_ON_USAGE ∑-permute
176+
"Warning: ∑-permute was deprecated in v1.5.
177+
Please use ∑-permute from `Algebra.Properties.CommutativeMonoid.Summation` instead."
178+
#-}
151179

152180
-- If the function takes the same value at 'i' and 'j', then transposing 'i' and
153181
-- 'j' then selecting 'j' is the same as selecting 'i'.

0 commit comments

Comments
 (0)