@@ -52,41 +52,35 @@ open CommutativeMonoid M
52
52
open import Algebra.Definitions _≈_
53
53
open import Relation.Binary.Reasoning.Setoid setoid
54
54
55
- module _ {n} where
56
- open B.Setoid (TE.setoid setoid n) public
57
- using ()
58
- renaming (_≈_ to _≋_)
59
55
60
- -- When summing over a function from a finite set, we can pull out any value and move it to the front.
61
56
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 _≋_)
75
66
76
- -- '_≈_' is a congruence over 'sumTable n'.
67
+ -- Version 1.5
77
68
78
69
sumₜ-cong-≈ : ∀ {n} → sumₜ {n} Preserves _≋_ ⟶ _≈_
79
70
sumₜ-cong-≈ {zero} p = refl
80
71
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
+ #-}
83
76
84
77
sumₜ-cong-≡ : ∀ {n} → sumₜ {n} Preserves _≗_ ⟶ _≡_
85
78
sumₜ-cong-≡ {zero} p = P.refl
86
79
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
+ #-}
90
84
91
85
sumₜ-idem-replicate : ∀ n {x} → _+_ IdempotentOn x → sumₜ (replicate {n = suc n} x) ≈ x
92
86
sumₜ-idem-replicate zero idem = +-identityʳ _
@@ -95,16 +89,38 @@ sumₜ-idem-replicate (suc n) {x} idem = begin
95
89
(x + x) + sumₜ (replicate {n = n} x) ≈⟨ +-congʳ idem ⟩
96
90
x + sumₜ (replicate {n = n} x) ≈⟨ sumₜ-idem-replicate n idem ⟩
97
91
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
+ #-}
100
96
101
97
sumₜ-zero : ∀ n → sumₜ (replicate {n = n} 0#) ≈ 0#
102
98
sumₜ-zero n = begin
103
99
sumₜ (replicate {n = n} 0#) ≈⟨ sym (+-identityˡ _) ⟩
104
100
0# + sumₜ (replicate {n = n} 0#) ≈⟨ sumₜ-idem-replicate n (+-identityˡ 0#) ⟩
105
101
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
+ #-}
106
106
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
+ #-}
108
124
109
125
∑-distrib-+ : ∀ n (f g : Fin n → Carrier) → ∑[ i < n ] (f i + g i) ≈ ∑[ i < n ] f i + ∑[ i < n ] g i
110
126
∑-distrib-+ zero f g = sym (+-identityˡ _)
@@ -119,17 +135,21 @@ sumₜ-zero n = begin
119
135
∑f = ∑[ i < n ] f (suc i)
120
136
∑g = ∑[ i < n ] g (suc i)
121
137
∑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
+ #-}
124
142
125
143
∑-comm : ∀ n m (f : Fin n → Fin m → Carrier) → ∑[ i < n ] ∑[ j < m ] f i j ≈ ∑[ j < m ] ∑[ i < n ] f i j
126
144
∑-comm zero m f = sym (sumₜ-zero m)
127
145
∑-comm (suc n) m f = begin
128
146
∑[ j < m ] f zero j + ∑[ i < n ] ∑[ j < m ] f (suc i) j ≈⟨ +-congˡ (∑-comm n m _) ⟩
129
147
∑[ j < m ] f zero j + ∑[ j < m ] ∑[ i < n ] f (suc i) j ≈⟨ sym (∑-distrib-+ m _ _) ⟩
130
148
∑[ 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
+ #-}
133
153
134
154
sumₜ-permute : ∀ {m n} t (π : Permutation m n) → sumₜ t ≈ sumₜ (permute π t)
135
155
sumₜ-permute {zero} {zero} t π = refl
@@ -145,9 +165,17 @@ sumₜ-permute {suc m} {suc n} t π = begin
145
165
where
146
166
0i = zero
147
167
π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
+ #-}
148
172
149
173
∑-permute : ∀ {m n} f (π : Permutation m n) → ∑[ i < n ] f i ≈ ∑[ i < m ] f (π ⟨$⟩ʳ i)
150
174
∑-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
+ #-}
151
179
152
180
-- If the function takes the same value at 'i' and 'j', then transposing 'i' and
153
181
-- 'j' then selecting 'j' is the same as selecting 'i'.
0 commit comments