9
9
open import Algebra.Bundles using (CommutativeSemigroup)
10
10
11
11
module Algebra.Properties.CommutativeSemigroup
12
- {a ℓ} (CS : CommutativeSemigroup a ℓ)
12
+ {a ℓ} (commutativeSemigroup : CommutativeSemigroup a ℓ)
13
13
where
14
14
15
- open CommutativeSemigroup CS
15
+ open CommutativeSemigroup commutativeSemigroup
16
16
open import Algebra.Definitions _≈_
17
17
open import Relation.Binary.Reasoning.Setoid setoid
18
18
open import Data.Product.Base using (_,_)
@@ -25,8 +25,8 @@ open import Algebra.Properties.Semigroup semigroup public
25
25
------------------------------------------------------------------------
26
26
-- Properties
27
27
28
- interchange : Interchangable _∙_ _∙_
29
- interchange a b c d = begin
28
+ medial : Medial _∙_
29
+ medial a b c d = begin
30
30
(a ∙ b) ∙ (c ∙ d) ≈⟨ assoc a b (c ∙ d) ⟩
31
31
a ∙ (b ∙ (c ∙ d)) ≈⟨ ∙-congˡ (assoc b c d) ⟨
32
32
a ∙ ((b ∙ c) ∙ d) ≈⟨ ∙-congˡ (∙-congʳ (comm b c)) ⟩
@@ -171,3 +171,18 @@ middleSemimedial x y z = begin
171
171
172
172
semimedial : Semimedial _∙_
173
173
semimedial = semimedialˡ , semimedialʳ
174
+
175
+
176
+ ------------------------------------------------------------------------
177
+ -- DEPRECATED NAMES
178
+ ------------------------------------------------------------------------
179
+ -- Please use the new names as continuing support for the old names is
180
+ -- not guaranteed.
181
+
182
+ -- Version 2.3
183
+
184
+ interchange = medial
185
+ {-# WARNING_ON_USAGE interchange
186
+ "Warning: interchange was deprecated in v2.3.
187
+ Please use medial instead."
188
+ #-}
0 commit comments